true -> false
{-
- Identity types, names and formulas
+ Path types, names and formulas
--------------------------------------------------------------------------
An element in ID A B is a line in the universe connecting A and B:
|- ID A B Type
-IdP is heterogeneous equality:
+PathP is heterogeneous equality:
|- P : ID A B |- a : A |- b : B
-----------------------------------------------
- |- IdP P a b Type
+ |- PathP P a b Type
-}
--- The usual identity can be seen a special case of IdP:
-Id (A : U) (a b : A) : U = IdP (<i> A) a b
+-- The constant Path type can be seen a special case of PathP:
+Path (A : U) (a b : A) : U = PathP (<i> A) a b
-- "<i>" abstracts the name/color/dimension i:
-refl (A : U) (a : A) : Id A a a = <i> a
+refl (A : U) (a : A) : Path A a a = <i> a
{-
-A proof "p : Id A a b" is thought of as a line between a and b:
+A proof "p : Path A a b" is thought of as a line between a and b:
p
a ------> b
-A proof "sq : Id (Id A a b) p q" is thought of as a square between p and q:
+A proof "sq : Path (Path A a b) p q" is thought of as a square between p and q:
q
-- It is possible to take the face of a path to obtain its endpoints:
-face0 (A : U) (a b : A) (p : Id A a b) : Id A a a = refl A (p @ 0)
-face1 (A : U) (a b : A) (p : Id A a b) : Id A b b = refl A (p @ 1)
+face0 (A : U) (a b : A) (p : Path A a b) : Path A a a = refl A (p @ 0)
+face1 (A : U) (a b : A) (p : Path A a b) : Path A b b = refl A (p @ 1)
-- By applying a path to a name i, "p @ i", it is seen as a path "in
-- dimension i" connecting "p @ 0" to "p @ 1". This way we get a
-- simple proof of cong:
-cong (A B : U) (f : A -> B) (a b : A) (p : Id A a b) : Id B (f a) (f b) =
+cong (A B : U) (f : A -> B) (a b : A) (p : Path A a b) : Path B (f a) (f b) =
<i> f (p @ i)
-- This also gives a short proof of function extensionality:
funExt (A : U) (B : A -> U) (f g : (x : A) -> B x)
- (p : (x : A) -> Id (B x) (f x) (g x)) : Id ((y : A) -> B y) f g =
+ (p : (x : A) -> Path (B x) (f x) (g x)) : Path ((y : A) -> B y) f g =
<i> \(x : A) -> (p x) @ i
{-
-}
-- Applying a path to a negated name inverts it:
-sym (A : U) (a b : A) (p : Id A a b) : Id A b a = <i> p @ -i
+sym (A : U) (a b : A) (p : Path A a b) : Path A b a = <i> p @ -i
-- This operation is an involution:
-symK (A : U) (a b : A) (p : Id A a b) : Id (Id A a b) p p =
- refl (Id A a b) (sym A b a (sym A a b p))
+symK (A : U) (a b : A) (p : Path A a b) : Path (Path A a b) p p =
+ refl (Path A a b) (sym A b a (sym A a b p))
{- Connections:
-- This gives a simple proof that singletons are contractible:
-singl (A : U) (a : A) : U = (x : A) * Id A a x
+singl (A : U) (a : A) : U = (x : A) * Path A a x
-contrSingl (A : U) (a b : A) (p : Id A a b) :
- Id (singl A a) (a,refl A a) (b,p) =
+contrSingl (A : U) (a b : A) (p : Path A a b) :
+ Path (singl A a) (a,refl A a) (b,p) =
<i> (p @ i, <j> p @ i /\ j)
Note that formulas does not form a boolean algebra, but a de Morgan
algebra. This means that "p @ 0" and "p @ i /\ -i" are different!
-testDeMorganAlgebra (A : U) (a b : A) (p : Id A a b) :
- Id (Id A a a) (<_> p @ 0) (<_> p @ 0) =
- refl (Id A a a) (<i> p @ i /\ -i)
+testDeMorganAlgebra (A : U) (a b : A) (p : Path A a b) :
+ Path (Path A a a) (<_> p @ 0) (<_> p @ 0) =
+ refl (Path A a a) (<i> p @ i /\ -i)
This is clear with the intuition that /\ correspond to min. More about
this later...
-}
-testConj (A : U) (a b : A) (p : Id A a b) : Id A a a = <i> p @ i /\ -i
-testDisj (A : U) (a b : A) (p : Id A a b) : Id A b b = <i> p @ i \/ -i
+testConj (A : U) (a b : A) (p : Path A a b) : Path A a a = <i> p @ i /\ -i
+testDisj (A : U) (a b : A) (p : Path A a b) : Path A b b = <i> p @ i \/ -i
{-
-So far we have: MLTT + Id, <i>, @, i/\j...
+So far we have: MLTT + Path, <i>, @, i/\j...
We can prove that singletons are contractible. But, we don't have
transport:
-transportf (A : U) (P : A -> U) (a b : A) (p : Id A a b) : P a -> P b
+transportf (A : U) (P : A -> U) (a b : A) (p : Path A a b) : P a -> P b
Which is what is needed to get J as:
side (opposite of the principal side).
-Transitivity of Id is obtained from a composition of this open square:
+Transitivity of Path is obtained from a composition of this open square:
i : I, j : I |-
p @ i
The composition computes the dashed line at the top of the square.
- ( <i>(<_>a)@i ) = (<j> (<_>a)@j) : Id (Id A a a) (<i> a) (<j> a)
+ ( <i>(<_>a)@i ) = (<j> (<_>a)@j) : Path (Path A a a) (<i> a) (<j> a)
p = <i> p @ i
-}
-pathscomp0 (A : U) (a b c : A) (p : Id A a b) (q : Id A b c) : Id A a c =
+pathscomp0 (A : U) (a b c : A) (p : Path A a b) (q : Path A b c) : Path A a c =
<i> comp (<_> A) (p @ i) [ (i = 0) -> <_> a
, (i = 1) -> q ]
-- The first two nonzero h-levels are propositions and sets:
-prop (A : U) : U = (a b : A) -> Id A a b
-set (A : U) : U = (a b : A) -> prop (Id A a b)
+prop (A : U) : U = (a b : A) -> Path A a b
+set (A : U) : U = (a b : A) -> prop (Path A a b)
-- Using compositions we get a short proof that any prop is a set. To
-- understand this proof one should draw an open cube with the back
-- face being a constant square in a and the sides given by the system
-- below.
propSet (A : U) (h : prop A) : set A =
- \(a b : A) (p q : Id A a b) ->
+ \(a b : A) (p q : Path A a b) ->
<i j> comp (<_> A) a [ (j=0) -> h a a
, (j=1) -> h a b
, (i=0) -> h a (p @ j)
v
-}
-Square (A : U) (a0 a1 b0 b1 : A) (u : Id A a0 a1) (v : Id A b0 b1)
- (r0 : Id A a0 b0) (r1 : Id A a1 b1) : U
- = IdP (<i> (IdP (<j> A) (u @ i) (v @ i))) r0 r1
+Square (A : U) (a0 a1 b0 b1 : A) (u : Path A a0 a1) (v : Path A b0 b1)
+ (r0 : Path A a0 b0) (r1 : Path A a1 b1) : U
+ = PathP (<i> (PathP (<j> A) (u @ i) (v @ i))) r0 r1
-constSquare (A : U) (a : A) (p : Id A a a) : Square A a a a a p p p p =
+constSquare (A : U) (a : A) (p : Path A a a) : Square A a a a a p p p p =
<i j> comp (<_> A) a [ (i = 0) -> <k> p @ j \/ - k
, (i = 1) -> <k> p @ j /\ k
, (j = 0) -> <k> p @ i \/ - k
-- have only considered constant lines..
-- "Kan composition":
-kan (A : U) (a b c d : A) (p : Id A a b)
- (q : Id A a c) (r : Id A b d) : Id A c d =
+kan (A : U) (a b c d : A) (p : Path A a b)
+ (q : Path A a c) (r : Path A b d) : Path A c d =
<i> comp (<_> A) (p @ i) [ (i = 0) -> q, (i = 1) -> r ]
-- Generalized Kan composition:
-kan' (A B : U) (P : Id U A B) (a b : A) (c d : B) (p : Id A a b)
- (q : IdP P a c) (r : IdP P b d) : Id B c d =
+kan' (A B : U) (P : Path U A B) (a b : A) (c d : B) (p : Path A a b)
+ (q : PathP P a c) (r : PathP P b d) : Path B c d =
<i> comp P (p @ i) [ (i = 0) -> q, (i = 1) -> r ]
-kan'' (A : U) (a b c d : A) (p : Id A a b)
- (q : Id A a c) (r : Id A b d) : Id A c d =
+kan'' (A : U) (a b c d : A) (p : Path A a b)
+ (q : Path A a c) (r : Path A b d) : Path A c d =
kan' A A (<_> A) a b c d p q r
suc b -> suc (add a b)
-- Exercise (for solution see bottom of the file):
-addZero : (a : nat) -> Id nat (add zero a) a = undefined
+addZero : (a : nat) -> Path nat (add zero a) a = undefined
-addSuc (a : nat) : (b : nat) -> Id nat (add (suc a) b) (suc (add a b)) = split
+addSuc (a : nat) : (b : nat) -> Path nat (add (suc a) b) (suc (add a b)) = split
zero -> <i> suc a
suc b' -> <i> suc (addSuc a b' @ i)
-- Exercise (can be done with a composition):
-addCom (a : nat) : (b : nat) -> Id nat (add a b) (add b a) = undefined
+addCom (a : nat) : (b : nat) -> Path nat (add a b) (add b a) = undefined
-addAssoc (a b : nat) : (c : nat) -> Id nat (add a (add b c)) (add (add a b) c) = split
+addAssoc (a b : nat) : (c : nat) -> Path nat (add a (add b c)) (add (add a b) c) = split
zero -> <i> add a b
suc c' -> <i> suc (addAssoc a b c' @ i)
Transport takes a path in the universe between A and B and produces a
function from A to B:
- transport : Id U A B -> A -> B
+ transport : Path U A B -> A -> B
-}
-- This gives a simple proof of subst (called transportf above and in
-- UniMath):
-subst (A : U) (P : A -> U) (a b : A) (p : Id A a b) (e : P a) : P b =
+subst (A : U) (P : A -> U) (a b : A) (p : Path A a b) (e : P a) : P b =
transport (cong A U P a b p) e
-- Transport is defined in the Haskell code as a composition with an
-- empty cube/system. It can also be directly defined as:
-trans (A B : U) (p : Id U A B) (a : A) : B = comp p a []
+trans (A B : U) (p : Path U A B) (a : A) : B = comp p a []
-- However these are not exactly the same because the lack of implicit
-- arguments.
-- subst with the fact that singletons are contractible this gives
-- the J eliminator:
-J (A : U) (a : A) (C : (x : A) -> Id A a x -> U)
- (d : C a (refl A a)) (x : A) (p : Id A a x) : C x p =
+J (A : U) (a : A) (C : (x : A) -> Path A a x -> U)
+ (d : C a (refl A a)) (x : A) (p : Path A a x) : C x p =
subst (singl A a) T (a,refl A a) (x,p) (contrSingl A a x p) d
where T (bp : singl A a) : U = C bp.1 bp.2
-- Note: Transporting with refl is not the identity function:
--- transRefl (A : U) (a : A) : Id A a a = refl A (transport (refl U A) a)
+-- transRefl (A : U) (a : A) : Path A a a = refl A (transport (refl U A) a)
-- This implies that the elimination rule for J does not hold definitonally:
--- defEqJ (A : U) (a : A) (C : (x : A) -> Id A a x -> U) (d : C a (refl A a)) :
--- Id (C a (refl A a)) (J A a C d a (refl A a)) d = refl (C a (refl A a)) d
+-- defEqJ (A : U) (a : A) (C : (x : A) -> Path A a x -> U) (d : C a (refl A a)) :
+-- Path (C a (refl A a)) (J A a C d a (refl A a)) d = refl (C a (refl A a)) d
-- We can get the equality between a and the transport of a by filling:
-transFill (A : U) (a : A) : Id A a (transport (<_> A) a) =
+transFill (A : U) (a : A) : Path A a (transport (<_> A) a) =
fill (<_> A) a []
-- composition in the universe).
-- A "strange" term showing what happens without regularity
--- strange (A : U) (a b c : A) (p : Id A a b) (q : Id A b c) :
--- Id A (comp (<_> A) a []) c =
+-- strange (A : U) (a b c : A) (p : Path A a b) (q : Path A b c) :
+-- Path A (comp (<_> A) a []) c =
-- <i> comp (<_> A) (p @ i) [ (i = 1) -> q ]
--- It seems like we can recover a new Id type with the correct
+-- It seems like we can recover a new Path type with the correct
-- definitional equalities using ideas from awfs of Andrew
-- Swan... Simon has implemented this in the branch "defeq".
{-
-So far we have: MLTT + Id, <i>, @, i/\j, compositions, transport
+So far we have: MLTT + Path, <i>, @, i/\j, compositions, transport
With all of this we get J, but so far without def. eq.
-}
-isContr (A : U) : U = (x : A) * ((y : A) -> Id A x y)
+isContr (A : U) : U = (x : A) * ((y : A) -> Path A x y)
fiber (A B : U) (f : A -> B) (y : B) : U =
- (x : A) * Id B y (f x)
+ (x : A) * Path B y (f x)
isEquiv (A B : U) (f : A -> B) : U =
(y : B) -> isContr (fiber A B f y)
-- Using the grad lemma we can transform isomorphisms into equalities.
lemIso (A B : U) (f : A -> B) (g : B -> A)
- (s : (y : B) -> Id B (f (g y)) y)
- (t : (x : A) -> Id A (g (f x)) x)
- (y : B) (x0 x1 : A) (p0 : Id B y (f x0)) (p1 : Id B y (f x1)) :
- Id (fiber A B f y) (x0,p0) (x1,p1) = <i> (p @ i,sq1 @ i)
+ (s : (y : B) -> Path B (f (g y)) y)
+ (t : (x : A) -> Path A (g (f x)) x)
+ (y : B) (x0 x1 : A) (p0 : Path B y (f x0)) (p1 : Path B y (f x1)) :
+ Path (fiber A B f y) (x0,p0) (x1,p1) = <i> (p @ i,sq1 @ i)
where
- rem0 : Id A (g y) x0 =
+ rem0 : Path A (g y) x0 =
<i> comp (<_> A) (g (p0 @ i)) [ (i = 1) -> t x0, (i = 0) -> <_> g y ]
- rem1 : Id A (g y) x1 =
+ rem1 : Path A (g y) x1 =
<i> comp (<_> A) (g (p1 @ i)) [ (i = 1) -> t x1, (i = 0) -> <_> g y ]
- p : Id A x0 x1 =
+ p : Path A x0 x1 =
<i> comp (<_> A) (g y) [ (i = 0) -> rem0
, (i = 1) -> rem1 ]
, (j = 0) -> s y ]
gradLemma (A B : U) (f : A -> B) (g : B -> A)
- (s : (y : B) -> Id B (f (g y)) y)
- (t : (x : A) -> Id A (g (f x)) x) : isEquiv A B f =
+ (s : (y : B) -> Path B (f (g y)) y)
+ (t : (x : A) -> Path A (g (f x)) x) : isEquiv A B f =
\ (y:B) -> ((g y,<i>s y@-i),\ (z:fiber A B f y) -> lemIso A B f g s t y (g y) z.1 (<i>s y@-i) z.2)
-isoId (A B : U) (f : A -> B) (g : B -> A)
- (s : (y : B) -> Id B (f (g y)) y)
- (t : (x : A) -> Id A (g (f x)) x) : Id U A B =
+isoPath (A B : U) (f : A -> B) (g : B -> A)
+ (s : (y : B) -> Path B (f (g y)) y)
+ (t : (x : A) -> Path A (g (f x)) x) : Path U A B =
<i> glue B [ (i = 0) -> (A,f,gradLemma A B f g s t)
, (i = 1) -> (B,idfun B,idIsEquiv B) ]
---------------------------------------------------------
-- Not is involutive:
-notK : (b : bool) -> Id bool (not (not b)) b = split
+notK : (b : bool) -> Path bool (not (not b)) b = split
false -> <i> false
true -> <i> true
-- This defines a non-trivial equality between bool and bool:
-notEq : Id U bool bool = isoId bool bool not not notK notK
+notEq : Path U bool bool = isoPath bool bool not not notK notK
-- We can transport true along this non-trivial equality:
testFalse : bool = transport notEq true
zero -> inl zero
suc n -> inr n
-sucpredZ : (x : Z) -> Id Z (sucZ (predZ x)) x = split
+sucpredZ : (x : Z) -> Path Z (sucZ (predZ x)) x = split
inl u -> <i> inl u
inr v -> lem v
where
- lem : (u : nat) -> Id Z (sucZ (predZ (inr u))) (inr u) = split
+ lem : (u : nat) -> Path Z (sucZ (predZ (inr u))) (inr u) = split
zero -> <i> inr zero
suc n -> <i> inr (suc n)
-predsucZ : (x : Z) -> Id Z (predZ (sucZ x)) x = split
+predsucZ : (x : Z) -> Path Z (predZ (sucZ x)) x = split
inl u -> lem u
where
- lem : (u : nat) -> Id Z (predZ (sucZ (inl u))) (inl u) = split
+ lem : (u : nat) -> Path Z (predZ (sucZ (inl u))) (inl u) = split
zero -> <i> inl zero
suc n -> <i> inl (suc n)
inr v -> <i> inr v
-sucIdZ : Id U Z Z = isoId Z Z sucZ predZ sucpredZ predsucZ
+sucPathZ : Path U Z Z = isoPath Z Z sucZ predZ sucpredZ predsucZ
-- We can transport along the proof forward and backwards:
-testOneZ : Z = transport sucIdZ zeroZ
-testNOneZ : Z = transport (<i> sucIdZ @ - i) zeroZ
+testOneZ : Z = transport sucPathZ zeroZ
+testNOneZ : Z = transport (<i> sucPathZ @ - i) zeroZ
--- transport sucIdZ = sucZ ?
+-- transport sucPathZ = sucZ ?
, (i = 1) -> neg zero ]
-- Nice version of the zero constructor:
-zeroPath : Id int (pos zero) (neg zero) = <i> zeroP {int} @ i
+zeroPath : Path int (pos zero) (neg zero) = <i> zeroP {int} @ i
sucInt : int -> int = split
pos n -> pos (suc n)
inl n -> neg (suc n)
inr n -> pos n
-toZK : (a : Z) -> Id Z (toZ (fromZ a)) a = split
+toZK : (a : Z) -> Path Z (toZ (fromZ a)) a = split
inl n -> <i> inl n
inr n -> <i> inr n
-fromZK : (a : int) -> Id int (fromZ (toZ a)) a = split
+fromZK : (a : int) -> Path int (fromZ (toZ a)) a = split
pos n -> <i> pos n
neg n -> rem n
- where rem : (n : nat) -> Id int (fromZ (toZ (neg n))) (neg n) = split
+ where rem : (n : nat) -> Path int (fromZ (toZ (neg n))) (neg n) = split
zero -> zeroPath
suc m -> <i> neg (suc m)
zeroP @ i -> <j> zeroPath @ i /\ j -- A connection makes this proof easy!
-isoIntZ : Id U Z int = isoId Z int fromZ toZ fromZK toZK
+isoIntZ : Path U Z int = isoPath Z int fromZ toZ fromZK toZK
data S1 = base
| loop <i> [ (i = 0) -> base, (i = 1) -> base]
-loopS1 : U = Id S1 base base
+loopS1 : U = Path S1 base base
-- The loop constructor
loop1 : loopS1 = <i> loop{S1} @ i
helix : S1 -> U = split
base -> Z
- loop @ i -> sucIdZ @ i
+ loop @ i -> sucPathZ @ i
-- The winding number:
winding (p : loopS1) : Z = transport (<i> helix (p @ i)) zeroZ
loopZN1 : Z = winding invLoop
loopZ0 : Z = winding (compS1 loop1 invLoop)
-mLoop : (x : S1) -> Id S1 x x = split
+mLoop : (x : S1) -> Path S1 x x = split
base -> loop1
loop @ i -> constSquare S1 base loop1 @ i
-- A nice example of a homotopy on the circle. The path going halfway
-- around the circle and then back is contractible:
-hmtpy : Id loopS1 (<i> base) (<i> loop{S1} @ (i /\ -i)) =
+hmtpy : Path loopS1 (<i> base) (<i> loop{S1} @ (i /\ -i)) =
<j i> loop{S1} @ j /\ i /\ -i
{- Solutions:
-addZero : (a : nat) -> Id nat (add zero a) a = split
+addZero : (a : nat) -> Path nat (add zero a) a = split
zero -> <i> zero
suc a' -> <i> suc (addZero a' @ i)
-addCom (a : nat) : (b : nat) -> Id nat (add a b) (add b a) = split
+addCom (a : nat) : (b : nat) -> Path nat (add a b) (add b a) = split
zero -> <i> addZero a @ -i
suc b' -> <i> comp (<_> nat) (suc (addCom a b' @ i)) [ (i = 0) -> <j> suc (add a b')
, (i = 1) -> <j> addSuc b' a @ -j ]
x1 p -> hS (x0 p) (posInd (\(p : pos) -> P (x0 p)) (hS pos1 h1) H p)
in f p
-sucPosSuc : (p : pos) -> Id nat (PosToN (sucPos p)) (suc (PosToN p)) = split
+sucPosSuc : (p : pos) -> Path nat (PosToN (sucPos p)) (suc (PosToN p)) = split
pos1 -> <_> suc (suc zero)
x0 p -> <_> suc (doubleN (PosToN p))
x1 p -> <i> doubleN (sucPosSuc p @ i)
-zeronPosToN (p : pos) : neg (Id nat zero (PosToN p)) =
- posInd (\(p : pos) -> neg (Id nat zero (PosToN p))) (\(prf : Id nat zero one) -> znots zero prf) hS p
+zeronPosToN (p : pos) : neg (Path nat zero (PosToN p)) =
+ posInd (\(p : pos) -> neg (Path nat zero (PosToN p))) (\(prf : Path nat zero one) -> znots zero prf) hS p
where
- hS (p : pos) (_ : neg (Id nat zero (PosToN p))) (prf : Id nat zero (PosToN (sucPos p))) : N0 =
+ hS (p : pos) (_ : neg (Path nat zero (PosToN p))) (prf : Path nat zero (PosToN (sucPos p))) : N0 =
znots (PosToN p) (<i> comp (<j> nat) (prf @ i) [ (i=0) -> <_> zero, (i = 1) -> sucPosSuc p ])
-- Inverse of PosToN:
zero -> pos1
suc n -> NtoPos' n
-PosToNK : (n : nat) -> Id nat (PosToN (NtoPos (suc n))) (suc n) = split
+PosToNK : (n : nat) -> Path nat (PosToN (NtoPos (suc n))) (suc n) = split
zero -> <_> (suc zero)
suc n ->
- let ih : Id nat (suc (PosToN (NtoPos' n))) (suc (suc n)) = <i> suc (PosToNK n @ i)
- in compId nat (PosToN (NtoPos (suc (suc n)))) (suc (PosToN (NtoPos' n)))
+ let ih : Path nat (suc (PosToN (NtoPos' n))) (suc (suc n)) = <i> suc (PosToNK n @ i)
+ in compPath nat (PosToN (NtoPos (suc (suc n)))) (suc (PosToN (NtoPos' n)))
(suc (suc n)) (sucPosSuc (NtoPos' n)) ih
-NtoPosSuc : (n : nat) -> neg (Id nat zero n) -> Id pos (NtoPos (suc n)) (sucPos (NtoPos n)) = split
- zero -> \(p : neg (Id nat zero zero)) -> efq (Id pos (NtoPos (suc zero)) (sucPos (NtoPos zero))) (p (<_> zero))
- suc n -> \(_ : neg (Id nat zero (suc n))) -> <_> (sucPos (NtoPos' n))
+NtoPosSuc : (n : nat) -> neg (Path nat zero n) -> Path pos (NtoPos (suc n)) (sucPos (NtoPos n)) = split
+ zero -> \(p : neg (Path nat zero zero)) -> efq (Path pos (NtoPos (suc zero)) (sucPos (NtoPos zero))) (p (<_> zero))
+ suc n -> \(_ : neg (Path nat zero (suc n))) -> <_> (sucPos (NtoPos' n))
-NtoPosK (p:pos) : Id pos (NtoPos (PosToN p)) p
- = posInd (\(p:pos) -> Id pos (NtoPos (PosToN p)) p) (refl pos pos1) hS p
+NtoPosK (p:pos) : Path pos (NtoPos (PosToN p)) p
+ = posInd (\(p:pos) -> Path pos (NtoPos (PosToN p)) p) (refl pos pos1) hS p
where
- hS (p : pos) (hp: Id pos (NtoPos (PosToN p)) p) : Id pos (NtoPos (PosToN (sucPos p))) (sucPos p)
+ hS (p : pos) (hp: Path pos (NtoPos (PosToN p)) p) : Path pos (NtoPos (PosToN (sucPos p))) (sucPos p)
=
- let H1 : Id pos (NtoPos (PosToN (sucPos p))) (NtoPos (suc (PosToN p)))
+ let H1 : Path pos (NtoPos (PosToN (sucPos p))) (NtoPos (suc (PosToN p)))
= mapOnPath nat pos NtoPos (PosToN (sucPos p)) (suc (PosToN p)) (sucPosSuc p)
- H2 : Id pos (NtoPos (suc (PosToN p))) (sucPos (NtoPos (PosToN p)))
+ H2 : Path pos (NtoPos (suc (PosToN p))) (sucPos (NtoPos (PosToN p)))
= NtoPosSuc (PosToN p) (zeronPosToN p)
- H3 : Id pos (sucPos (NtoPos (PosToN p))) (sucPos p)
+ H3 : Path pos (sucPos (NtoPos (PosToN p))) (sucPos p)
= mapOnPath pos pos sucPos (NtoPos (PosToN p)) p hp
- in compId pos (NtoPos (PosToN (sucPos p))) (sucPos (NtoPos (PosToN p))) (sucPos p)
- (compId pos (NtoPos (PosToN (sucPos p))) (NtoPos (suc (PosToN p))) (sucPos (NtoPos (PosToN p))) H1 H2)
+ in compPath pos (NtoPos (PosToN (sucPos p))) (sucPos (NtoPos (PosToN p))) (sucPos p)
+ (compPath pos (NtoPos (PosToN (sucPos p))) (NtoPos (suc (PosToN p))) (sucPos (NtoPos (PosToN p))) H1 H2)
H3
posToNinj : injective pos nat PosToN
- = \ (p0 p1:pos) (h:Id nat (PosToN p0) (PosToN p1)) ->
+ = \ (p0 p1:pos) (h:Path nat (PosToN p0) (PosToN p1)) ->
<i> comp (<_>pos) (NtoPos (h@i)) [(i=0) -> NtoPosK p0,(i=1) -> NtoPosK p1]
-- Binary natural numbers
binN0 -> zero
binNpos p -> PosToN p
-NtoBinNK : (n:nat) -> Id nat (BinNtoN (NtoBinN n)) n = split
+NtoBinNK : (n:nat) -> Path nat (BinNtoN (NtoBinN n)) n = split
zero -> refl nat zero
suc n -> PosToNK n
-rem (p : pos) : Id binN (NtoBinN (PosToN (sucPos p))) (binNpos (sucPos p)) =
- compId binN (NtoBinN (PosToN (sucPos p))) (binNpos (NtoPos (suc (PosToN p))))
+rem (p : pos) : Path binN (NtoBinN (PosToN (sucPos p))) (binNpos (sucPos p)) =
+ compPath binN (NtoBinN (PosToN (sucPos p))) (binNpos (NtoPos (suc (PosToN p))))
(binNpos (sucPos p)) rem1 rem2
where
- rem1 : Id binN (NtoBinN (PosToN (sucPos p))) (binNpos (NtoPos (suc (PosToN p))))
+ rem1 : Path binN (NtoBinN (PosToN (sucPos p))) (binNpos (NtoPos (suc (PosToN p))))
= mapOnPath nat binN NtoBinN (PosToN (sucPos p)) (suc (PosToN p)) (sucPosSuc p)
- rem2 : Id binN (binNpos (NtoPos (suc (PosToN p)))) (binNpos (sucPos p))
+ rem2 : Path binN (binNpos (NtoPos (suc (PosToN p)))) (binNpos (sucPos p))
= <i>binNpos
- (compId pos (NtoPos (suc (PosToN p))) (sucPos (NtoPos (PosToN p))) (sucPos p)
+ (compPath pos (NtoPos (suc (PosToN p))) (sucPos (NtoPos (PosToN p))) (sucPos p)
(NtoPosSuc (PosToN p) (zeronPosToN p))
(mapOnPath pos pos sucPos (NtoPos (PosToN p)) p (NtoPosK p))@i)
-lem1 (p : pos) : Id binN (NtoBinN (PosToN p)) (binNpos p)
- = posInd (\ (p:pos) -> Id binN (NtoBinN (PosToN p)) (binNpos p)) (refl binN (binNpos pos1))
- (\ (p:pos) (_:Id binN (NtoBinN (PosToN p)) (binNpos p)) -> rem p) p
+lem1 (p : pos) : Path binN (NtoBinN (PosToN p)) (binNpos p)
+ = posInd (\ (p:pos) -> Path binN (NtoBinN (PosToN p)) (binNpos p)) (refl binN (binNpos pos1))
+ (\ (p:pos) (_:Path binN (NtoBinN (PosToN p)) (binNpos p)) -> rem p) p
-BinNtoNK : (b:binN) -> Id binN (NtoBinN (BinNtoN b)) b = split -- retract binN N BinNtoN NtoBinN = split
+BinNtoNK : (b:binN) -> Path binN (NtoBinN (BinNtoN b)) b = split -- retract binN N BinNtoN NtoBinN = split
binN0 -> refl binN binN0
binNpos p -> lem1 p
-IdbinNN : Id U binN nat
- = isoId binN nat BinNtoN NtoBinN NtoBinNK BinNtoNK
+PathbinNN : Path U binN nat
+ = isoPath binN nat BinNtoN NtoBinN NtoBinNK BinNtoNK
-- Doubling
-- 2^10 * e = 2^5 * (2^5 * e)
propDouble (D : Double) : U
- = Id (carrier D) (iter (carrier D) (doubleN five) (double D) (elt D))
+ = Path (carrier D) (iter (carrier D) (doubleN five) (double D) (elt D))
(iter (carrier D) five (double D) (iter (carrier D) five (double D) (elt D)))
-- The property we want to prove that takes long to typecheck:
DoubleBinN' : Double
= D binN doubleBinN' (NtoBinN n1024)
-eqDouble1 : Id Double DoubleN DoubleBinN'
- = elimIsIso nat (\(B : U) (f : nat -> B) (g : B -> nat) -> Id Double DoubleN (D B (\(b : B) -> f (doubleN (g b))) (f n1024)))
+eqDouble1 : Path Double DoubleN DoubleBinN'
+ = elimIsIso nat (\(B : U) (f : nat -> B) (g : B -> nat) -> Path Double DoubleN (D B (\(b : B) -> f (doubleN (g b))) (f n1024)))
(refl Double DoubleN) binN NtoBinN BinNtoN BinNtoNK NtoBinNK
-eqDouble2 : Id Double DoubleBinN' DoubleBinN
+eqDouble2 : Path Double DoubleBinN' DoubleBinN
= mapOnPath (binN -> binN) Double F doubleBinN' doubleBinN rem
where
F (d:binN -> binN) : Double
= D binN d (NtoBinN n1024)
- rem : Id (binN -> binN) doubleBinN' doubleBinN
+ rem : Path (binN -> binN) doubleBinN' doubleBinN
= funExt binN (\(x:binN) -> binN) doubleBinN' doubleBinN rem1
where
- rem1 : (n : binN) -> Id binN (doubleBinN' n) (doubleBinN n)
+ rem1 : (n : binN) -> Path binN (doubleBinN' n) (doubleBinN n)
= split
binN0 -> refl binN binN0
binNpos p ->
- let p1 : Id binN (NtoBinN (doubleN (PosToN p))) (NtoBinN (PosToN (x0 p)))
+ let p1 : Path binN (NtoBinN (doubleN (PosToN p))) (NtoBinN (PosToN (x0 p)))
= mapOnPath nat binN NtoBinN (doubleN (PosToN p)) (PosToN (x0 p)) (refl nat (doubleN (PosToN p)))
- in compId binN (NtoBinN (doubleN (PosToN p))) (NtoBinN (PosToN (x0 p))) (binNpos (x0 p)) p1 (lem1 (x0 p))
+ in compPath binN (NtoBinN (doubleN (PosToN p))) (NtoBinN (PosToN (x0 p))) (binNpos (x0 p)) p1 (lem1 (x0 p))
-eqDouble : Id Double DoubleN DoubleBinN
- = compId Double DoubleN DoubleBinN' DoubleBinN eqDouble1 eqDouble2
+eqDouble : Path Double DoubleN DoubleBinN
+ = compPath Double DoubleN DoubleBinN' DoubleBinN eqDouble1 eqDouble2
-- opaque doubleN
false -> f
true -> t
-falseNeqTrue : neg (Id bool false true) =
- \(h : Id bool false true) -> subst bool (caseBool U bool N0) false true h false
+falseNeqTrue : neg (Path bool false true) =
+ \(h : Path bool false true) -> subst bool (caseBool U bool N0) false true h false
-trueNeqFalse : neg (Id bool true false) =
- \(h : Id bool true false) -> subst bool (caseBool U N0 bool) true false h true
+trueNeqFalse : neg (Path bool true false) =
+ \(h : Path bool true false) -> subst bool (caseBool U N0 bool) true false h true
-boolDec : (b1 b2 : bool) -> dec (Id bool b1 b2) = split
+boolDec : (b1 b2 : bool) -> dec (Path bool b1 b2) = split
false -> rem
where
- rem : (b : bool) -> dec (Id bool false b) = split
+ rem : (b : bool) -> dec (Path bool false b) = split
false -> inl (<i> false)
true -> inr falseNeqTrue
true -> rem
where
- rem : (b : bool) -> dec (Id bool true b) = split
+ rem : (b : bool) -> dec (Path bool true b) = split
false -> inr trueNeqFalse
true -> inl (<i> true)
setbool' : set bool = hedberg bool boolDec
-- Direct proof that bool is a set:
-lem1 : (y:bool) (p:Id bool true y) -> Id bool true y = split
- false -> \ (p : Id bool true false) -> p
- true -> \ (p : Id bool true true) -> <i>true
+lem1 : (y:bool) (p:Path bool true y) -> Path bool true y = split
+ false -> \ (p : Path bool true false) -> p
+ true -> \ (p : Path bool true true) -> <i>true
-lem2 : (x y :bool) (p:Id bool true x) (q:Id bool true y) -> Id bool x y = split
- false -> \ (y:bool) (p:Id bool true false) (q:Id bool true y) -> efq (Id bool false y) (trueNeqFalse p)
- true -> \ (y:bool) (p:Id bool true true) (q:Id bool true y) -> lem1 y q
+lem2 : (x y :bool) (p:Path bool true x) (q:Path bool true y) -> Path bool x y = split
+ false -> \ (y:bool) (p:Path bool true false) (q:Path bool true y) -> efq (Path bool false y) (trueNeqFalse p)
+ true -> \ (y:bool) (p:Path bool true true) (q:Path bool true y) -> lem1 y q
-lem3 : prop (Id bool true true) =
- \ (p q:Id bool true true) -> <j i>lem2 (p@i) (q@i) (<k>p@i/\k) (<k>q@i/\k) @ j
+lem3 : prop (Path bool true true) =
+ \ (p q:Path bool true true) -> <j i>lem2 (p@i) (q@i) (<k>p@i/\k) (<k>q@i/\k) @ j
-lem4 : (y:bool) (p:Id bool false y) -> Id bool false y = split
- false -> \ (p : Id bool false false) -> <i>false
- true -> \ (p : Id bool false true) -> p
+lem4 : (y:bool) (p:Path bool false y) -> Path bool false y = split
+ false -> \ (p : Path bool false false) -> <i>false
+ true -> \ (p : Path bool false true) -> p
-lem5 : (x y :bool) (p:Id bool false x) (q:Id bool false y) -> Id bool x y = split
- false -> \ (y:bool) (p:Id bool false false) (q:Id bool false y) -> lem4 y q
- true -> \ (y:bool) (p:Id bool false true) (q:Id bool false y) -> efq (Id bool true y) (falseNeqTrue p)
+lem5 : (x y :bool) (p:Path bool false x) (q:Path bool false y) -> Path bool x y = split
+ false -> \ (y:bool) (p:Path bool false false) (q:Path bool false y) -> lem4 y q
+ true -> \ (y:bool) (p:Path bool false true) (q:Path bool false y) -> efq (Path bool true y) (falseNeqTrue p)
-lem6 : prop (Id bool false false) =
- \ (p q:Id bool false false) -> <j i>lem5 (p@i) (q@i) (<k>p@i/\k) (<k>q@i/\k) @ j
+lem6 : prop (Path bool false false) =
+ \ (p q:Path bool false false) -> <j i>lem5 (p@i) (q@i) (<k>p@i/\k) (<k>q@i/\k) @ j
-lem7 : (y:bool) (p:Id bool false y) (q:Id bool false y) -> Id (Id bool false y) p q = split
+lem7 : (y:bool) (p:Path bool false y) (q:Path bool false y) -> Path (Path bool false y) p q = split
false -> lem6
- true -> \ (p:Id bool false true) (q:Id bool false true) -> efq (Id (Id bool false true) p q) (falseNeqTrue p)
+ true -> \ (p:Path bool false true) (q:Path bool false true) -> efq (Path (Path bool false true) p q) (falseNeqTrue p)
-lem8 : (y:bool) (p:Id bool true y) (q:Id bool true y) -> Id (Id bool true y) p q = split
- false -> \ (p:Id bool true false) (q:Id bool true false) -> efq (Id (Id bool true false) p q) (trueNeqFalse p)
+lem8 : (y:bool) (p:Path bool true y) (q:Path bool true y) -> Path (Path bool true y) p q = split
+ false -> \ (p:Path bool true false) (q:Path bool true false) -> efq (Path (Path bool true false) p q) (trueNeqFalse p)
true -> lem3
setbool : set bool = split
true -> false
-- negBool is involutive:
-negBoolK : (b : bool) -> Id bool (negBool (negBool b)) b = split
+negBoolK : (b : bool) -> Path bool (negBool (negBool b)) b = split
false -> refl bool false
true -> refl bool true
-- This defines a non-trivial equality between bool and bool:
-negBoolEq : Id U bool bool =
- isoId bool bool negBool negBool negBoolK negBoolK
+negBoolEq : Path U bool bool =
+ isoPath bool bool negBool negBool negBoolK negBoolK
-- We can transport true along this non-trivial equality:
testFalse : bool = transport negBoolEq true
false -> zeroF2
true -> oneF2
-f2ToBoolK : (x : F2) -> Id F2 (boolToF2 (f2ToBool x)) x = split
+f2ToBoolK : (x : F2) -> Path F2 (boolToF2 (f2ToBool x)) x = split
zeroF2 -> refl F2 zeroF2
oneF2 -> refl F2 oneF2
-boolToF2K : (b : bool) -> Id bool (f2ToBool (boolToF2 b)) b = split
+boolToF2K : (b : bool) -> Path bool (f2ToBool (boolToF2 b)) b = split
false -> refl bool false
true -> refl bool true
-boolEqF2 : Id U bool F2 =
- isoId bool F2 boolToF2 f2ToBool f2ToBoolK boolToF2K
+boolEqF2 : Path U bool F2 =
+ isoPath bool F2 boolToF2 f2ToBool f2ToBoolK boolToF2K
negF2 : F2 -> F2 = subst U (\(X : U) -> (X -> X)) bool F2 boolEqF2 negBool
--- lemTest (A : U) : (B : U) (p : Id U A B) (a : A) -> IdP p a (transport p a) =
--- J U A (\(B : U) (p : Id U A B) -> (a : A) -> IdP p a (transport p a)) (refl A)
+-- lemTest (A : U) : (B : U) (p : Path U A B) (a : A) -> PathP p a (transport p a) =
+-- J U A (\(B : U) (p : Path U A B) -> (a : A) -> PathP p a (transport p a)) (refl A)
--- test : IdP boolEqF2 true oneF2 = <i> glueElem oneF2 [ (i = 0) -> true ]
--- test1 : IdP boolEqF2 true oneF2 = lemTest bool F2 boolEqF2 true
+-- test : PathP boolEqF2 true oneF2 = <i> glueElem oneF2 [ (i = 0) -> true ]
+-- test1 : PathP boolEqF2 true oneF2 = lemTest bool F2 boolEqF2 true
-F2EqBool : Id U F2 bool = inv U bool F2 boolEqF2
+F2EqBool : Path U F2 bool = inv U bool F2 boolEqF2
negBool' : bool -> bool = subst U (\(X : U) -> (X -> X)) F2 bool F2EqBool negF2
-F2EqBoolComp : Id U F2 bool =
- compId U F2 bool bool F2EqBool negBoolEq
+F2EqBoolComp : Path U F2 bool =
+ compPath U F2 bool bool F2EqBool negBoolEq
test2 : bool = trans F2 bool F2EqBoolComp oneF2
-negNegEq : Id U bool bool =
- compId U bool bool bool negBoolEq negBoolEq
+negNegEq : Path U bool bool =
+ compPath U bool bool bool negBoolEq negBoolEq
test3 : bool = trans bool bool negNegEq true
-test4 : Id U bool bool = <i> negNegEq @ i
+test4 : Path U bool bool = <i> negNegEq @ i
-kanBool : Id U bool bool =
+kanBool : Path U bool bool =
kan U bool bool bool bool negBoolEq negBoolEq negBoolEq
squareBoolF2 : Square U bool bool bool F2 (refl U bool)
boolEqF2 (refl U bool) boolEqF2 =
<i j> boolEqF2 @ i /\ j
-test5 : IdP boolEqF2 true oneF2 =
+test5 : PathP boolEqF2 true oneF2 =
<i> comp (<j> boolEqF2 @ i /\ j) true []
-test6 : Id bool true true =
+test6 : Path bool true true =
<i> comp (<j> F2EqBool @ i \/ j) (test5 @ - i) []
-test7 : Id U F2 F2 =
- subst U (\(X:U) -> Id U X X) bool F2 boolEqF2 negNegEq
+test7 : Path U F2 F2 =
+ subst U (\(X:U) -> Path U X X) bool F2 boolEqF2 negNegEq
-test8 : Id U F2 F2 =
- subst U (\(X:U) -> Id U X X) bool F2 boolEqF2 (refl U bool)
+test8 : Path U F2 F2 =
+ subst U (\(X:U) -> Path U X X) bool F2 boolEqF2 (refl U bool)
-test9 : Id U F2 F2 =
- comp (<i> Id U (boolEqF2 @ i) (boolEqF2 @ i)) (refl U bool) []
+test9 : Path U F2 F2 =
+ comp (<i> Path U (boolEqF2 @ i) (boolEqF2 @ i)) (refl U bool) []
-p : Id U F2 bool = <i> comp (<_> U) bool [ (i = 0) -> boolEqF2, (i = 1) -> <_> bool]
-q : Id U F2 F2 = <i> p @ (i /\ - i)
+p : Path U F2 bool = <i> comp (<_> U) bool [ (i = 0) -> boolEqF2, (i = 1) -> <_> bool]
+q : Path U F2 F2 = <i> p @ (i /\ - i)
-- A small test of univalence using boolean negation
isEquivNegBool : isEquiv bool bool negBool =
gradLemma bool bool negBool negBool negBoolK negBoolK
-eqBool : Id U bool bool =
- trans (equiv bool bool) (Id U bool bool)
+eqBool : Path U bool bool =
+ trans (equiv bool bool) (Path U bool bool)
(<i> corrUniv bool bool @ -i) (negBool,isEquivNegBool)
testf : bool = trans bool bool eqBool true
testt : bool = trans bool bool eqBool false
-- Some tests of normal forms of univalence:
-testUniv1 (A : U) : Id U A A =
- trans (equiv A A) (Id U A A) (<i> corrUniv A A @ -i) (idEquiv A)
+testUniv1 (A : U) : Path U A A =
+ trans (equiv A A) (Path U A A) (<i> corrUniv A A @ -i) (idEquiv A)
-- obtained by normal form
-ntestUniv1 (A:U) : Id U A A =
+ntestUniv1 (A:U) : Path U A A =
<i> comp (<_>U)
(comp (<_>U) A
- [ (i = 0) -> <l> comp (<_>U) A [ (l = 0) -> <i> glue A [ (i = 0) -> (A,(\(x : A) -> x,\(a : A) -> ((a,<l> a),\(z : ((x : A) * IdP (<l> A) a x)) -> <l> (z.2 @ l,<i> z.2 @ (l /\ i))))), (i = 1) -> (A,(comp (<l> A -> A) (\(x : A) -> x) [],comp (<l> (y : A) ->
- ((x : ((x : A) * IdP (<j> A) y (comp (<j> A) (comp (<i> A) x []) [ (l = 0) -> <j> comp (<k> A) x [ (j = 1) -> <k> x ] ]))) * (y0 : ((x0 : A) * IdP (<j> A) y (comp (<j> A) (comp (<i> A) x0 []) [ (l = 0) -> <j> comp (<k> A) x0 [ (j = 1) -> <k> x0 ] ]))) -> IdP (<j> ((x0 : A) * IdP (<j> A) y (comp (<j> A) (comp (<i> A) x0 []) [ (l = 0) -> <j> comp (<k> A) x0 [ (j = 1) -> <k> x0 ] ]))) x y0)) (\(a : A) -> ((a,<l> a),\(z : ((x : A) * IdP (<l> A) a x)) -> <l> (z.2 @ l,<i> z.2 @ (l /\ i)))) [])) ], (l = 1) -> <i> glue A [ (i = 0) -> (A,(\(x : A) -> x,\(a : A) -> ((a,<l> a),\(z : ((x : A) * IdP (<l> A) a x)) -> <l> (z.2 @ l,<i> z.2 @ (l /\ i))))), (i = 1) -> (A,(comp (<l> A -> A) (\(x : A) -> x) [],comp (<l> (y : A) ->
- ((x : ((x : A) * IdP (<j> A) y (comp (<j> A) (comp (<i> A) x []) [ (l = 0) -> <j> comp (<k> A) x [ (j = 1) -> <k> x ] ]))) * (y0 : ((x0 : A) * IdP (<j> A) y (comp (<j> A) (comp (<i> A) x0 []) [ (l = 0) -> <j> comp (<k> A) x0 [ (j = 1) -> <k> x0 ] ]))) -> IdP (<j> ((x0 : A) * IdP (<j> A) y (comp (<j> A) (comp (<i> A) x0 []) [ (l = 0) -> <j> comp (<k> A) x0 [ (j = 1) -> <k> x0 ] ]))) x y0)) (\(a : A) -> ((a,<l> a),\(z : ((x : A) * IdP (<l> A) a x)) -> <l> (z.2 @ l,<i> z.2 @ (l /\ i)))) [])) ] ], (i = 1) -> <l> A ]) [ (i = 0) -> <l> A, (i = 1) -> <l> A ]
+ [ (i = 0) -> <l> comp (<_>U) A [ (l = 0) -> <i> glue A [ (i = 0) -> (A,(\(x : A) -> x,\(a : A) -> ((a,<l> a),\(z : ((x : A) * PathP (<l> A) a x)) -> <l> (z.2 @ l,<i> z.2 @ (l /\ i))))), (i = 1) -> (A,(comp (<l> A -> A) (\(x : A) -> x) [],comp (<l> (y : A) ->
+ ((x : ((x : A) * PathP (<j> A) y (comp (<j> A) (comp (<i> A) x []) [ (l = 0) -> <j> comp (<k> A) x [ (j = 1) -> <k> x ] ]))) * (y0 : ((x0 : A) * PathP (<j> A) y (comp (<j> A) (comp (<i> A) x0 []) [ (l = 0) -> <j> comp (<k> A) x0 [ (j = 1) -> <k> x0 ] ]))) -> PathP (<j> ((x0 : A) * PathP (<j> A) y (comp (<j> A) (comp (<i> A) x0 []) [ (l = 0) -> <j> comp (<k> A) x0 [ (j = 1) -> <k> x0 ] ]))) x y0)) (\(a : A) -> ((a,<l> a),\(z : ((x : A) * PathP (<l> A) a x)) -> <l> (z.2 @ l,<i> z.2 @ (l /\ i)))) [])) ], (l = 1) -> <i> glue A [ (i = 0) -> (A,(\(x : A) -> x,\(a : A) -> ((a,<l> a),\(z : ((x : A) * PathP (<l> A) a x)) -> <l> (z.2 @ l,<i> z.2 @ (l /\ i))))), (i = 1) -> (A,(comp (<l> A -> A) (\(x : A) -> x) [],comp (<l> (y : A) ->
+ ((x : ((x : A) * PathP (<j> A) y (comp (<j> A) (comp (<i> A) x []) [ (l = 0) -> <j> comp (<k> A) x [ (j = 1) -> <k> x ] ]))) * (y0 : ((x0 : A) * PathP (<j> A) y (comp (<j> A) (comp (<i> A) x0 []) [ (l = 0) -> <j> comp (<k> A) x0 [ (j = 1) -> <k> x0 ] ]))) -> PathP (<j> ((x0 : A) * PathP (<j> A) y (comp (<j> A) (comp (<i> A) x0 []) [ (l = 0) -> <j> comp (<k> A) x0 [ (j = 1) -> <k> x0 ] ]))) x y0)) (\(a : A) -> ((a,<l> a),\(z : ((x : A) * PathP (<l> A) a x)) -> <l> (z.2 @ l,<i> z.2 @ (l /\ i)))) [])) ] ], (i = 1) -> <l> A ]) [ (i = 0) -> <l> A, (i = 1) -> <l> A ]
testUniv2 : bool = trans bool bool (ntestUniv1 bool) true
ntestUniv2 : bool =
- comp (<i> comp (<_>U) (comp (<_>U) bool [ (i = 0) -> <j> comp (<_>U) bool [ (j = 0) -> <i> glue bool [ (i = 0) -> (bool,(\(x : bool) -> x,\(a : bool) -> ((a,<i> a),\(z : ((x : bool) * IdP (<i> bool) a x)) -> <i> (z.2 @ i,<j> z.2 @ (i /\ j))))), (i = 1) -> (bool,(comp (<i> bool -> bool) (\(x : bool) -> x) [],comp (<i> (y : bool) ->
- ((x : ((x : bool) * IdP (<j> bool) y (comp (<j> bool) (comp (<i> bool) x []) [ (i = 0) -> <j> comp (<k> bool) x [ (j = 1) -> <k> x ] ]))) * (y0 : ((x0 : bool) * IdP (<j> bool) y (comp (<j> bool) (comp (<i> bool) x0 []) [ (i = 0) -> <j> comp (<k> bool) x0 [ (j = 1) -> <k> x0 ] ]))) -> IdP (<j> ((x0 : bool) * IdP (<k> bool) y (comp (<j> bool) (comp (<i> bool) x0 []) [ (i = 0) -> <j> comp (<k> bool) x0 [ (j = 1) -> <k> x0 ] ]))) x y0)) (\(a : bool) -> ((a,<i> a),\(z : ((x : bool) * IdP (<i> bool) a x)) -> <i> (z.2 @ i,<j> z.2 @ (i /\ j)))) [])) ], (j = 1) -> <i> glue bool [ (i = 0) -> (bool,(\(x : bool) -> x,\(a : bool) -> ((a,<i> a),\(z : ((x : bool) * IdP (<i> bool) a x)) -> <i> (z.2 @ i,<j> z.2 @ (i /\ j))))), (i = 1) -> (bool,(comp (<i> bool -> bool) (\(x : bool) -> x) [],comp (<i> (y : bool) ->
- ((x : ((x : bool) * IdP (<j> bool) y (comp (<j> bool) (comp (<i> bool) x []) [ (i = 0) -> <j> comp (<k> bool) x [ (j = 1) -> <k> x ] ]))) * (y0 : ((x0 : bool) * IdP (<j> bool) y (comp (<j> bool) (comp (<i> bool) x0 []) [ (i = 0) -> <j> comp (<k> bool) x0 [ (j = 1) -> <k> x0 ] ]))) -> IdP (<j> ((x0 : bool) * IdP (<k> bool) y (comp (<j> bool) (comp (<i> bool) x0 []) [ (i = 0) -> <j> comp (<k> bool) x0 [ (j = 1) -> <k> x0 ] ]))) x y0)) (\(a : bool) -> ((a,<i> a),\(z : ((x : bool) * IdP (<i> bool) a x)) -> <i> (z.2 @ i,<j> z.2 @ (i /\ j)))) [])) ] ], (i = 1) -> <j> bool ]) [ (i = 0) -> <j> bool, (i = 1) -> <j> bool ]) true []
+ comp (<i> comp (<_>U) (comp (<_>U) bool [ (i = 0) -> <j> comp (<_>U) bool [ (j = 0) -> <i> glue bool [ (i = 0) -> (bool,(\(x : bool) -> x,\(a : bool) -> ((a,<i> a),\(z : ((x : bool) * PathP (<i> bool) a x)) -> <i> (z.2 @ i,<j> z.2 @ (i /\ j))))), (i = 1) -> (bool,(comp (<i> bool -> bool) (\(x : bool) -> x) [],comp (<i> (y : bool) ->
+ ((x : ((x : bool) * PathP (<j> bool) y (comp (<j> bool) (comp (<i> bool) x []) [ (i = 0) -> <j> comp (<k> bool) x [ (j = 1) -> <k> x ] ]))) * (y0 : ((x0 : bool) * PathP (<j> bool) y (comp (<j> bool) (comp (<i> bool) x0 []) [ (i = 0) -> <j> comp (<k> bool) x0 [ (j = 1) -> <k> x0 ] ]))) -> PathP (<j> ((x0 : bool) * PathP (<k> bool) y (comp (<j> bool) (comp (<i> bool) x0 []) [ (i = 0) -> <j> comp (<k> bool) x0 [ (j = 1) -> <k> x0 ] ]))) x y0)) (\(a : bool) -> ((a,<i> a),\(z : ((x : bool) * PathP (<i> bool) a x)) -> <i> (z.2 @ i,<j> z.2 @ (i /\ j)))) [])) ], (j = 1) -> <i> glue bool [ (i = 0) -> (bool,(\(x : bool) -> x,\(a : bool) -> ((a,<i> a),\(z : ((x : bool) * PathP (<i> bool) a x)) -> <i> (z.2 @ i,<j> z.2 @ (i /\ j))))), (i = 1) -> (bool,(comp (<i> bool -> bool) (\(x : bool) -> x) [],comp (<i> (y : bool) ->
+ ((x : ((x : bool) * PathP (<j> bool) y (comp (<j> bool) (comp (<i> bool) x []) [ (i = 0) -> <j> comp (<k> bool) x [ (j = 1) -> <k> x ] ]))) * (y0 : ((x0 : bool) * PathP (<j> bool) y (comp (<j> bool) (comp (<i> bool) x0 []) [ (i = 0) -> <j> comp (<k> bool) x0 [ (j = 1) -> <k> x0 ] ]))) -> PathP (<j> ((x0 : bool) * PathP (<k> bool) y (comp (<j> bool) (comp (<i> bool) x0 []) [ (i = 0) -> <j> comp (<k> bool) x0 [ (j = 1) -> <k> x0 ] ]))) x y0)) (\(a : bool) -> ((a,<i> a),\(z : ((x : bool) * PathP (<i> bool) a x)) -> <i> (z.2 @ i,<j> z.2 @ (i /\ j)))) [])) ] ], (i = 1) -> <j> bool ]) [ (i = 0) -> <j> bool, (i = 1) -> <j> bool ]) true []
import nat
import univalence
-lemReflComp (A : U) (a b : A) (p : Id A a b) : Id (Id A a b) (compId A a a b (<_> a) p) p =
+lemReflComp (A : U) (a b : A) (p : Path A a b) : Path (Path A a b) (compPath A a a b (<_> a) p) p =
<j i> comp (<k> A) (p @ i /\ j) [(i=0) -> <_> a, (j=1) -> <_> p @ i, (i=1) -> <k> p @ k \/ j ]
opaque lemReflComp
-lemReflComp' (A : U) (a b : A) (p : Id A a b) : Id (Id A a b) (compId A a b b p (<_> b)) p =
+lemReflComp' (A : U) (a b : A) (p : Path A a b) : Path (Path A a b) (compPath A a b b p (<_> b)) p =
<j i> comp (<k> A) (p @ i) [(i=0) -> <_> a, (j=1) -> <_> p @ i, (i=1) -> <_> b ]
opaque lemReflComp'
setPi (A : U) (B : A -> U) (h : (x : A) -> set (B x))
(f0 f1 : (x : A) -> B x)
- (p1 p2 : Id ((x : A) -> B x) f0 f1)
- : Id (Id ((x : A) -> B x) f0 f1) p1 p2
+ (p1 p2 : Path ((x : A) -> B x) f0 f1)
+ : Path (Path ((x : A) -> B x) f0 f1) p1 p2
= <i j> \(x : A) -> (h x (f0 x) (f1 x) (<i> (p1@i) x) (<i> (p2@i) x)) @ i @ j
opaque setPi
-lemIdPProp (A B : U) (AProp : prop A) (p : Id U A B) : (x : A) -> (y : B) -> IdP p x y
- = J U A (\(B : U) -> \(p : Id U A B) -> (x : A) -> (y : B) -> IdP p x y) AProp B p
-opaque lemIdPProp
+lemPathPProp (A B : U) (AProp : prop A) (p : Path U A B) : (x : A) -> (y : B) -> PathP p x y
+ = J U A (\(B : U) -> \(p : Path U A B) -> (x : A) -> (y : B) -> PathP p x y) AProp B p
+opaque lemPathPProp
-lemIdPSet (A B : U) (ASet : set A) (p : Id U A B) : (x : A) (y : B) (s t : IdP p x y) -> Id (IdP p x y) s t
- = J U A (\(B : U) -> \(p : Id U A B) -> (x : A) (y : B) (s t : IdP p x y) -> Id (IdP p x y) s t) ASet B p
-opaque lemIdPSet
+lemPathPSet (A B : U) (ASet : set A) (p : Path U A B) : (x : A) (y : B) (s t : PathP p x y) -> Path (PathP p x y) s t
+ = J U A (\(B : U) -> \(p : Path U A B) -> (x : A) (y : B) (s t : PathP p x y) -> Path (PathP p x y) s t) ASet B p
+opaque lemPathPSet
-lemIdPSet2 (A B : U) (ASet : set A) (p1 : Id U A B)
- : (p2 : Id U A B) -> (p : Id (Id U A B) p1 p2) ->
- (x : A) -> (y : B) -> (s : IdP p1 x y) -> (t : IdP p2 x y) -> IdP (<i> (IdP (p @ i) x y)) s t
- = J (Id U A B) p1 (\(p2 : Id U A B) -> \(p : Id (Id U A B) p1 p2) -> (x : A) -> (y : B) -> (s : IdP p1 x y) -> (t : IdP p2 x y) -> IdP (<i> (IdP (p @ i) x y)) s t)
- (lemIdPSet A B ASet p1)
-opaque lemIdPSet2
+lemPathPSet2 (A B : U) (ASet : set A) (p1 : Path U A B)
+ : (p2 : Path U A B) -> (p : Path (Path U A B) p1 p2) ->
+ (x : A) -> (y : B) -> (s : PathP p1 x y) -> (t : PathP p2 x y) -> PathP (<i> (PathP (p @ i) x y)) s t
+ = J (Path U A B) p1 (\(p2 : Path U A B) -> \(p : Path (Path U A B) p1 p2) -> (x : A) -> (y : B) -> (s : PathP p1 x y) -> (t : PathP p2 x y) -> PathP (<i> (PathP (p @ i) x y)) s t)
+ (lemPathPSet A B ASet p1)
+opaque lemPathPSet2
-substIdP (A B : U) (p : Id U A B) (x : A) (y : B) (q : Id B (transport p x) y) : IdP p x y
- = transport (<i> IdP p x (q@i)) hole
+substPathP (A B : U) (p : Path U A B) (x : A) (y : B) (q : Path B (transport p x) y) : PathP p x y
+ = transport (<i> PathP p x (q@i)) hole
where
- hole : IdP p x (transport p x) = <i> comp (<j> p @ (i /\ j)) x [(i=0) -> <_> x]
-opaque substIdP
+ hole : PathP p x (transport p x) = <i> comp (<j> p @ (i /\ j)) x [(i=0) -> <_> x]
+opaque substPathP
-transRefl (A : U) (a : A) : Id A (transport (<_> A) a) a = <i> comp (<_> A) a [(i=1) -> <_>a]
+transRefl (A : U) (a : A) : Path A (transport (<_> A) a) a = <i> comp (<_> A) a [(i=1) -> <_>a]
opaque transRefl
-isContrProp (A : U) (p : isContr A) (x y : A) : Id A x y = compId A x p.1 y (<i> p.2 x @ -i) (p.2 y)
+isContrProp (A : U) (p : isContr A) (x y : A) : Path A x y = compPath A x p.1 y (<i> p.2 x @ -i) (p.2 y)
opaque isContrProp
equivProp (A B : U) (AProp : prop A) (BProp : prop B) (F : A -> B) (G : B -> A) : equiv A B =
(F, gradLemma A B F G (\(y : B) -> BProp (F (G y)) y) (\(x : A) -> AProp (G (F x)) x))
opaque equivProp
-idProp (A B : U) (AProp : prop A) (BProp : prop B) (F : A -> B) (G : B -> A) : Id U A B
- = equivId A B (equivProp A B AProp BProp F G).1 (equivProp A B AProp BProp F G).2
+idProp (A B : U) (AProp : prop A) (BProp : prop B) (F : A -> B) (G : B -> A) : Path U A B
+ = equivPath A B (equivProp A B AProp BProp F G).1 (equivProp A B AProp BProp F G).2
opaque idProp
-lemTransEquiv (A B:U) (e:Id U A B) (x:A) : Id B (transport e x) ((transport (<i> equiv (e@-i) B) (idEquiv B)).1 x)
+lemTransEquiv (A B:U) (e:Path U A B) (x:A) : Path B (transport e x) ((transport (<i> equiv (e@-i) B) (idEquiv B)).1 x)
= <i> transRefl B (transport e x)@-i
opaque lemTransEquiv
= let A : U = C.1
hom : A -> A -> U = C.2
in (homSet : (x y : A) -> set (hom x y))
- * (cIdL : (x y : A) -> (f : hom x y) -> Id (hom x y) (c x x y (id x) f) f)
- * (cIdR : (x y : A) -> (f : hom x y) -> Id (hom x y) (c x y y f (id y)) f)
- * ( (x y z w : A) -> (f : hom x y) -> (g : hom y z) -> (h : hom z w) -> Id (hom x w) (c x z w (c x y z f g) h) (c x y w f (c y z w g h)))
+ * (cPathL : (x y : A) -> (f : hom x y) -> Path (hom x y) (c x x y (id x) f) f)
+ * (cPathR : (x y : A) -> (f : hom x y) -> Path (hom x y) (c x y y f (id y)) f)
+ * ( (x y z w : A) -> (f : hom x y) -> (g : hom y z) -> (h : hom z w) -> Path (hom x w) (c x z w (c x y z f g) h) (c x y w f (c y z w g h)))
propIsPrecategory2 (C : categoryData) (id : (x : C.1) -> C.2 x x) (c : (x y z : C.1) -> C.2 x y -> C.2 y z -> C.2 x z)
: prop (isPrecategory2 C id c)
propSig
((x y : A) -> set (hom x y))
(\(_:(x y : A) -> set (hom x y))->
- ((cIdL : (x y : A) -> (f : hom x y) -> Id (hom x y) (c x x y (id x) f) f)
- * (cIdR : (x y : A) -> (f : hom x y) -> Id (hom x y) (c x y y f (id y)) f)
- * ( (x y z w : A) -> (f : hom x y) -> (g : hom y z) -> (h : hom z w) -> Id (hom x w) (c x z w (c x y z f g) h) (c x y w f (c y z w g h)))))
+ ((cPathL : (x y : A) -> (f : hom x y) -> Path (hom x y) (c x x y (id x) f) f)
+ * (cPathR : (x y : A) -> (f : hom x y) -> Path (hom x y) (c x y y f (id y)) f)
+ * ( (x y z w : A) -> (f : hom x y) -> (g : hom y z) -> (h : hom z w) -> Path (hom x w) (c x z w (c x y z f g) h) (c x y w f (c y z w g h)))))
(propPi A (\(x:A)->(y : A) -> set (hom x y))
(\(x:A)->propPi A (\(y : A) -> set (hom x y)) (\(y:A)->setIsProp (hom x y))))
(\(hset:(x y : A) -> set (hom x y))->
(propAnd
- ((x y : A) -> (f : hom x y) -> Id (hom x y) (c x x y (id x) f) f)
- ((cIdR : (x y : A) -> (f : hom x y) -> Id (hom x y) (c x y y f (id y)) f)
- * ( (x y z w : A) -> (f : hom x y) -> (g : hom y z) -> (h : hom z w) -> Id (hom x w) (c x z w (c x y z f g) h) (c x y w f (c y z w g h))))
- (propPi A (\(x:A)->(y : A) -> (f : hom x y) -> Id (hom x y) (c x x y (id x) f) f)
- (\(x:A)->propPi A (\(y : A) -> (f : hom x y) -> Id (hom x y) (c x x y (id x) f) f)
- (\(y:A)->propPi (hom x y) (\(f : hom x y) -> Id (hom x y) (c x x y (id x) f) f)
+ ((x y : A) -> (f : hom x y) -> Path (hom x y) (c x x y (id x) f) f)
+ ((cPathR : (x y : A) -> (f : hom x y) -> Path (hom x y) (c x y y f (id y)) f)
+ * ( (x y z w : A) -> (f : hom x y) -> (g : hom y z) -> (h : hom z w) -> Path (hom x w) (c x z w (c x y z f g) h) (c x y w f (c y z w g h))))
+ (propPi A (\(x:A)->(y : A) -> (f : hom x y) -> Path (hom x y) (c x x y (id x) f) f)
+ (\(x:A)->propPi A (\(y : A) -> (f : hom x y) -> Path (hom x y) (c x x y (id x) f) f)
+ (\(y:A)->propPi (hom x y) (\(f : hom x y) -> Path (hom x y) (c x x y (id x) f) f)
(\(f:hom x y)->hset x y (c x x y (id x) f) f))))
(propAnd
- ((x y : A) -> (f : hom x y) -> Id (hom x y) (c x y y f (id y)) f)
- ((x y z w : A) -> (f : hom x y) -> (g : hom y z) -> (h : hom z w) -> Id (hom x w) (c x z w (c x y z f g) h) (c x y w f (c y z w g h)))
- (propPi A (\(x:A)->(y : A) -> (f : hom x y) -> Id (hom x y) (c x y y f (id y)) f)
- (\(x:A)->propPi A (\(y : A) -> (f : hom x y) -> Id (hom x y) (c x y y f (id y)) f)
- (\(y:A)->propPi (hom x y) (\(f : hom x y) -> Id (hom x y) (c x y y f (id y)) f)
+ ((x y : A) -> (f : hom x y) -> Path (hom x y) (c x y y f (id y)) f)
+ ((x y z w : A) -> (f : hom x y) -> (g : hom y z) -> (h : hom z w) -> Path (hom x w) (c x z w (c x y z f g) h) (c x y w f (c y z w g h)))
+ (propPi A (\(x:A)->(y : A) -> (f : hom x y) -> Path (hom x y) (c x y y f (id y)) f)
+ (\(x:A)->propPi A (\(y : A) -> (f : hom x y) -> Path (hom x y) (c x y y f (id y)) f)
+ (\(y:A)->propPi (hom x y) (\(f : hom x y) -> Path (hom x y) (c x y y f (id y)) f)
(\(f:hom x y)->hset x y (c x y y f (id y)) f))))
- (propPi A (\(x:A)->(y z w : A) -> (f : hom x y) -> (g : hom y z) -> (h : hom z w) -> Id (hom x w) (c x z w (c x y z f g) h) (c x y w f (c y z w g h)))
- (\(x:A)->propPi A (\(y:A)->(z w : A) -> (f : hom x y) -> (g : hom y z) -> (h : hom z w) -> Id (hom x w) (c x z w (c x y z f g) h) (c x y w f (c y z w g h)))
- (\(y:A)->propPi A (\(z:A)->(w : A) -> (f : hom x y) -> (g : hom y z) -> (h : hom z w) -> Id (hom x w) (c x z w (c x y z f g) h) (c x y w f (c y z w g h)))
- (\(z:A)->propPi A (\(w:A)-> (f : hom x y) -> (g : hom y z) -> (h : hom z w) -> Id (hom x w) (c x z w (c x y z f g) h) (c x y w f (c y z w g h)))
- (\(w:A)->propPi (hom x y) (\(f : hom x y) -> (g : hom y z) -> (h : hom z w) -> Id (hom x w) (c x z w (c x y z f g) h) (c x y w f (c y z w g h)))
- (\(f:hom x y)->propPi (hom y z) (\(g : hom y z) -> (h : hom z w) -> Id (hom x w) (c x z w (c x y z f g) h) (c x y w f (c y z w g h)))
- (\(g:hom y z)->propPi (hom z w) (\(h : hom z w) -> Id (hom x w) (c x z w (c x y z f g) h) (c x y w f (c y z w g h)))
+ (propPi A (\(x:A)->(y z w : A) -> (f : hom x y) -> (g : hom y z) -> (h : hom z w) -> Path (hom x w) (c x z w (c x y z f g) h) (c x y w f (c y z w g h)))
+ (\(x:A)->propPi A (\(y:A)->(z w : A) -> (f : hom x y) -> (g : hom y z) -> (h : hom z w) -> Path (hom x w) (c x z w (c x y z f g) h) (c x y w f (c y z w g h)))
+ (\(y:A)->propPi A (\(z:A)->(w : A) -> (f : hom x y) -> (g : hom y z) -> (h : hom z w) -> Path (hom x w) (c x z w (c x y z f g) h) (c x y w f (c y z w g h)))
+ (\(z:A)->propPi A (\(w:A)-> (f : hom x y) -> (g : hom y z) -> (h : hom z w) -> Path (hom x w) (c x z w (c x y z f g) h) (c x y w f (c y z w g h)))
+ (\(w:A)->propPi (hom x y) (\(f : hom x y) -> (g : hom y z) -> (h : hom z w) -> Path (hom x w) (c x z w (c x y z f g) h) (c x y w f (c y z w g h)))
+ (\(f:hom x y)->propPi (hom y z) (\(g : hom y z) -> (h : hom z w) -> Path (hom x w) (c x z w (c x y z f g) h) (c x y w f (c y z w g h)))
+ (\(g:hom y z)->propPi (hom z w) (\(h : hom z w) -> Path (hom x w) (c x z w (c x y z f g) h) (c x y w f (c y z w g h)))
(\(h:hom z w)->hset x w (c x z w (c x y z f g) h) (c x y w f (c y z w g h)))))))))))))
isPrecategory (C : categoryData) : U =
cH (C : precategory) (a b : cA C) : U = C.1.2 a b
cHSet (C : precategory) (a b : cA C) : set (cH C a b) = C.2.2.2.1 a b
cC (C : precategory) (x y z : cA C) (f : cH C x y) (g : cH C y z) : cH C x z = C.2.2.1 x y z f g
-cId (C : precategory) (x : cA C) : cH C x x = C.2.1 x
-cIdL (C : precategory) (x y : cA C) (f : cH C x y)
- : Id (cH C x y) (cC C x x y (cId C x) f) f = C.2.2.2.2.1 x y f
-cIdR (C : precategory) (x y : cA C) (f : cH C x y)
- : Id (cH C x y) (cC C x y y f (cId C y)) f = C.2.2.2.2.2.1 x y f
-cIdC (C : precategory) (x y z w : cA C) (f : cH C x y) (g : cH C y z) (h : cH C z w)
- : Id (cH C x w) (cC C x z w (cC C x y z f g) h) (cC C x y w f (cC C y z w g h)) = C.2.2.2.2.2.2 x y z w f g h
-
-catIso3 (A B : precategory) : U = (e1 : Id U (cA A) (cA B))
- * (e2 : IdP (<i>(x y : e1@i)->U) (cH A) (cH B))
- * (_ : IdP (<i> (x:e1@i)->(e2@i) x x) (cId A) (cId B))
- * (IdP (<i> (x y z:e1@i)->(e2@i) x y->(e2@i) y z->(e2@i) x z) (cC A) (cC B))
-
-eCatIso3 (A B : precategory) : Id U (Id precategory A B) (catIso3 A B)
- = isoId (Id precategory A B) (catIso3 A B)
+cPath (C : precategory) (x : cA C) : cH C x x = C.2.1 x
+cPathL (C : precategory) (x y : cA C) (f : cH C x y)
+ : Path (cH C x y) (cC C x x y (cPath C x) f) f = C.2.2.2.2.1 x y f
+cPathR (C : precategory) (x y : cA C) (f : cH C x y)
+ : Path (cH C x y) (cC C x y y f (cPath C y)) f = C.2.2.2.2.2.1 x y f
+cPathC (C : precategory) (x y z w : cA C) (f : cH C x y) (g : cH C y z) (h : cH C z w)
+ : Path (cH C x w) (cC C x z w (cC C x y z f g) h) (cC C x y w f (cC C y z w g h)) = C.2.2.2.2.2.2 x y z w f g h
+
+catIso3 (A B : precategory) : U = (e1 : Path U (cA A) (cA B))
+ * (e2 : PathP (<i>(x y : e1@i)->U) (cH A) (cH B))
+ * (_ : PathP (<i> (x:e1@i)->(e2@i) x x) (cPath A) (cPath B))
+ * (PathP (<i> (x y z:e1@i)->(e2@i) x y->(e2@i) y z->(e2@i) x z) (cC A) (cC B))
+
+eCatIso3 (A B : precategory) : Path U (Path precategory A B) (catIso3 A B)
+ = isoPath (Path precategory A B) (catIso3 A B)
F G FG GF
where
- F (e:Id precategory A B):catIso3 A B = (<i>(e@i).1.1, <i>(e@i).1.2, <i>(e@i).2.1, <i>(e@i).2.2.1)
- G (e:catIso3 A B):Id precategory A B = <i> ((e.1@i, e.2.1@i), e.2.2.1@i, e.2.2.2@i
- ,lemIdPProp (isPrecategory2 A.1 A.2.1 A.2.2.1)
+ F (e:Path precategory A B):catIso3 A B = (<i>(e@i).1.1, <i>(e@i).1.2, <i>(e@i).2.1, <i>(e@i).2.2.1)
+ G (e:catIso3 A B):Path precategory A B = <i> ((e.1@i, e.2.1@i), e.2.2.1@i, e.2.2.2@i
+ ,lemPathPProp (isPrecategory2 A.1 A.2.1 A.2.2.1)
(isPrecategory2 B.1 B.2.1 B.2.2.1)
(propIsPrecategory2 A.1 A.2.1 A.2.2.1)
(<i>isPrecategory2 (e.1@i, e.2.1@i) (e.2.2.1@i) (e.2.2.2@i))
A.2.2.2 B.2.2.2 @ i)
- FG(e:catIso3 A B):Id (catIso3 A B) (F (G e)) e=<_>e
- GF(e:Id precategory A B):Id (Id precategory A B) (G (F e)) e
+ FG(e:catIso3 A B):Path (catIso3 A B) (F (G e)) e=<_>e
+ GF(e:Path precategory A B):Path (Path precategory A B) (G (F e)) e
=<i j>((e@j).1, (e@j).2.1, (e@j).2.2.1
- ,lemIdPSet (isPrecategory2 A.1 A.2.1 A.2.2.1)
+ ,lemPathPSet (isPrecategory2 A.1 A.2.1 A.2.2.1)
(isPrecategory2 B.1 B.2.1 B.2.2.1)
(propSet (isPrecategory2 A.1 A.2.1 A.2.2.1) (propIsPrecategory2 A.1 A.2.1 A.2.2.1))
(<i>isPrecategory2 (e@i).1 (e@i).2.1 (e@i).2.2.1)
(<j> (G (F e)@j).2.2.2) (<j>(e@j).2.2.2)
@ i @ j)
-catIso30 (A B : precategory) : U = (e1 : Id U (cA A) (cA B))
- * (e2 : Id ((x y : cA A)->U) (cH A) (\(x y:cA A)->cH B (transport e1 x) (transport e1 y)))
- * (_ : Id ((x:cA A)->cH A x x)
- (cId A)
- (\(x : cA A) -> transport (<j>(e2@-j) x x) (cId B (transport e1 x))))
- * (Id ((x y z:cA A)->cH A x y->cH A y z->cH A x z)
+catIso30 (A B : precategory) : U = (e1 : Path U (cA A) (cA B))
+ * (e2 : Path ((x y : cA A)->U) (cH A) (\(x y:cA A)->cH B (transport e1 x) (transport e1 y)))
+ * (_ : Path ((x:cA A)->cH A x x)
+ (cPath A)
+ (\(x : cA A) -> transport (<j>(e2@-j) x x) (cPath B (transport e1 x))))
+ * (Path ((x y z:cA A)->cH A x y->cH A y z->cH A x z)
(\(x y z:cA A)(f:cH A x y)(g:cH A y z)->cC A x y z f g)
(\(x y z:cA A)(f:cH A x y)(g:cH A y z)->
transport (<i>(e2@-i) x z)
(transport (<i>(e2@i) x y) f)
(transport (<i>(e2@i) y z) g))))
-eCatIso30 (A B : precategory) : Id U (catIso3 A B) (catIso30 A B)
- = <i> (e1 : Id U (cA A) (cA B))
+eCatIso30 (A B : precategory) : Path U (catIso3 A B) (catIso30 A B)
+ = <i> (e1 : Path U (cA A) (cA B))
* (let e21 : (x y : e1@-i) -> U
= comp (<_> (x y : e1@-i) -> U)
(\(x y : e1@-i) -> cH B (transport (<j>e1@-i\/j) x) (transport (<j>e1@-i\/j) y))
[(i=1)-><_>\(x y:cA A) -> cH B (transport e1 x) (transport e1 y)
,(i=0)-><j>\(x y:cA B) -> cH B (transRefl (cA B) x @ j) (transRefl (cA B) y @ j)
]
- f21 : Id ((x y : e1@-i) -> U)
+ f21 : Path ((x y : e1@-i) -> U)
(\(x y : e1@-i) -> cH B (transport (<j>e1@-i\/j) x) (transport (<j>e1@-i\/j) y))
e21
= fill (<_> (x y : e1@-i) -> U)
,(i=0)-><j>\(x y:cA B) -> cH B (transRefl (cA B) x @ j) (transRefl (cA B) y @ j)
]
in
- (e2 : IdP (<j> (x y : e1@j/\-i) -> U) (cH A) e21)
- * (_ : IdP (<j> (x : e1@j/\-i) -> (e2@j/\-i) x x)
- (cId A)
+ (e2 : PathP (<j> (x y : e1@j/\-i) -> U) (cH A) e21)
+ * (_ : PathP (<j> (x : e1@j/\-i) -> (e2@j/\-i) x x)
+ (cPath A)
(comp (<_> (x:e1@-i)->(e2@-i) x x)
- (\(x : e1@-i) -> transport (<j>(e2@-i\/-j) x x) (transport (<j>(f21@j) x x) (cId B (transport (<j>e1@-i\/j) x))))
- [(i=1)-><j>\(x:cA A) -> transport (<j>(e2@-j) x x) (transRefl ((f21@0) x x) (cId B (transport e1 x)) @ j)
+ (\(x : e1@-i) -> transport (<j>(e2@-i\/-j) x x) (transport (<j>(f21@j) x x) (cPath B (transport (<j>e1@-i\/j) x))))
+ [(i=1)-><j>\(x:cA A) -> transport (<j>(e2@-j) x x) (transRefl ((f21@0) x x) (cPath B (transport e1 x)) @ j)
,(i=0)-><j>\(x:cA B) ->
transRefl ((e2@1) x x)
- (compId (cH B x x)
+ (compPath (cH B x x)
(transport (<k>cH B (transRefl (cA B) x@k) (transRefl (cA B) x@k))
- (cId B (transport (<_> cA B) x)))
- (transport (<_>cH B x x) (cId B x))
- (cId B x)
+ (cPath B (transport (<_> cA B) x)))
+ (transport (<_>cH B x x) (cPath B x))
+ (cPath B x)
(<j>transport (<k>cH B (transRefl (cA B) x@j\/k) (transRefl (cA B) x@j\/k))
- (cId B (transRefl (cA B) x @ j)))
- (transRefl (cH B x x) (cId B x))
+ (cPath B (transRefl (cA B) x @ j)))
+ (transRefl (cH B x x) (cPath B x))
@ j) @ j
]))
- * (IdP (<j> (x y z : e1@j/\-i) (f : (e2@j/\-i) x y) (g : (e2@j/\-i) y z) -> (e2@j/\-i) x z)
+ * (PathP (<j> (x y z : e1@j/\-i) (f : (e2@j/\-i) x y) (g : (e2@j/\-i) y z) -> (e2@j/\-i) x z)
(cC A)
(comp (<_> (x y z : e1@-i) (f : (e2@-i) x y) (g : (e2@-i) y z) -> (e2@-i) x z)
(\(x y z : e1@-i) (f : (e2@-i) x y) (g : (e2@-i) y z) ->
(transRefl ((f21@0) y z) (transport (<j>(e2@j) y z) g) @ j)) @ j)
,(i=0)-><j>\(x y z : cA B) (f : cH B x y) (g : cH B y z) ->
transRefl ((e2@1) x z)
- (compId (cH B x z)
+ (compPath (cH B x z)
(transport (<k>cH B (transRefl (cA B) x@k) (transRefl (cA B) z@k))
(cC B (transport (<_>cA B) x) (transport (<_>cA B) y) (transport (<_>cA B) z)
(transport (<k>cH B (transRefl (cA B) x@-k) (transRefl (cA B) y@-k)) (transport (<j>(e2@1) x y) f))
@ j) @ j
])))
-catIso311 (A B : precategory) (e1:Id U (cA A) (cA B)) : U
- = Id ((x y : cA A)->U) (cH A) (\(x y:cA A)->cH B (transport e1 x) (transport e1 y))
-catIso312 (A B : precategory) (e1:Id U (cA A) (cA B)) (e2:catIso311 A B e1) : U
- = Id ((x:cA A)->cH B (transport e1 x) (transport e1 x))
- (\(x:cA A)->transport (<i> (e2@i) x x) (cId A x))
- (\(x : cA A) -> cId B (transport e1 x))
-catIso313 (A B : precategory) (e1:Id U (cA A) (cA B)) (e2:catIso311 A B e1) : U
- = Id ((x y z:cA A)->cH A x y->cH A y z->cH B (transport e1 x) (transport e1 z))
+catIso311 (A B : precategory) (e1:Path U (cA A) (cA B)) : U
+ = Path ((x y : cA A)->U) (cH A) (\(x y:cA A)->cH B (transport e1 x) (transport e1 y))
+catIso312 (A B : precategory) (e1:Path U (cA A) (cA B)) (e2:catIso311 A B e1) : U
+ = Path ((x:cA A)->cH B (transport e1 x) (transport e1 x))
+ (\(x:cA A)->transport (<i> (e2@i) x x) (cPath A x))
+ (\(x : cA A) -> cPath B (transport e1 x))
+catIso313 (A B : precategory) (e1:Path U (cA A) (cA B)) (e2:catIso311 A B e1) : U
+ = Path ((x y z:cA A)->cH A x y->cH A y z->cH B (transport e1 x) (transport e1 z))
(\(x y z:cA A)(f:cH A x y)(g:cH A y z)->transport (<i>(e2@i)x z) (cC A x y z f g))
(\(x y z:cA A)(f:cH A x y)(g:cH A y z)->
cC B (transport e1 x) (transport e1 y) (transport e1 z)
(transport (<i>(e2@i) x y) f)
(transport (<i>(e2@i) y z) g))
-catIso31 (A B : precategory) : U = (e1 : Id U (cA A) (cA B))
+catIso31 (A B : precategory) : U = (e1 : Path U (cA A) (cA B))
* (e2 : catIso311 A B e1)
* (_ : catIso312 A B e1 e2)
* ( catIso313 A B e1 e2)
-eCatIso31 (A B : precategory) : Id U (catIso30 A B) (catIso31 A B)
- = <i> (e1 : Id U (cA A) (cA B))
- * (e2 : Id ((x y : cA A)->U) (cH A) (\(x y:cA A)->cH B (transport e1 x) (transport e1 y)))
+eCatIso31 (A B : precategory) : Path U (catIso30 A B) (catIso31 A B)
+ = <i> (e1 : Path U (cA A) (cA B))
+ * (e2 : Path ((x y : cA A)->U) (cH A) (\(x y:cA A)->cH B (transport e1 x) (transport e1 y)))
* (_:comp (<_> U)
- (Id ((x:cA A)->(e2@i) x x)
- (\(x:cA A)->transport (<j> (e2@i/\j) x x) (cId A x))
- (\(x:cA A)->transport (<j> (e2@i\/-j) x x) (cId B (transport e1 x))))
- [(i=0)-><k>Id ((x:cA A)->cH A x x)
- (\(x:cA A)->transRefl (cH A x x) (cId A x) @ k)
- (\(x:cA A)->transport (<j> (e2@-j) x x) (cId B (transport e1 x)))
- ,(i=1)-><k>Id ((x:cA A)->(e2@i) x x)
- (\(x:cA A)->transport (<j> (e2@j) x x) (cId A x))
- (\(x:cA A)->transRefl ((e2@1) x x) (cId B (transport e1 x)) @ k)
+ (Path ((x:cA A)->(e2@i) x x)
+ (\(x:cA A)->transport (<j> (e2@i/\j) x x) (cPath A x))
+ (\(x:cA A)->transport (<j> (e2@i\/-j) x x) (cPath B (transport e1 x))))
+ [(i=0)-><k>Path ((x:cA A)->cH A x x)
+ (\(x:cA A)->transRefl (cH A x x) (cPath A x) @ k)
+ (\(x:cA A)->transport (<j> (e2@-j) x x) (cPath B (transport e1 x)))
+ ,(i=1)-><k>Path ((x:cA A)->(e2@i) x x)
+ (\(x:cA A)->transport (<j> (e2@j) x x) (cPath A x))
+ (\(x:cA A)->transRefl ((e2@1) x x) (cPath B (transport e1 x)) @ k)
])
* (comp (<_> U)
- (Id ((x y z:cA A)->cH A x y->cH A y z->(e2@i) x z)
+ (Path ((x y z:cA A)->cH A x y->cH A y z->(e2@i) x z)
(\(x y z:cA A)(f:cH A x y)(g:cH A y z)->
transport (<j> (e2@i/\j) x z) (cC A x y z f g))
(\(x y z:cA A)(f:cH A x y)(g:cH A y z)->
(cC B (transport e1 x) (transport e1 y) (transport e1 z)
(transport (<j>(e2@j) x y) f)
(transport (<j>(e2@j) y z) g))))
- [(i=0)-><k>Id ((x y z:cA A)->cH A x y->cH A y z->cH A x z)
+ [(i=0)-><k>Path ((x y z:cA A)->cH A x y->cH A y z->cH A x z)
(\(x y z:cA A)(f:cH A x y)(g:cH A y z)->
transRefl ((e2@0) x z) (cC A x y z f g) @ k)
(\(x y z:cA A)(f:cH A x y)(g:cH A y z)->
(cC B (transport e1 x) (transport e1 y) (transport e1 z)
(transport (<j>(e2@j) x y) f)
(transport (<j>(e2@j) y z) g)))
- ,(i=1)-><k>Id ((x y z:cA A)->cH A x y->cH A y z->(e2@1) x z)
+ ,(i=1)-><k>Path ((x y z:cA A)->cH A x y->cH A y z->(e2@1) x z)
(\(x y z:cA A)(f:cH A x y)(g:cH A y z)->
transport (<j>(e2@j) x z) (cC A x y z f g))
(\(x y z:cA A)(f:cH A x y)(g:cH A y z)->
])
lemEquivCoh (A B:U) (e:equiv A B) (x:A)
- : Id (Id B (e.1 x) (e.1 (invEq A B e (e.1 x))))
+ : Path (Path B (e.1 x) (e.1 (invEq A B e (e.1 x))))
(<j>e.1(secEq A B e x@-j))
(<j>retEq A B e (e.1 x)@-j)
- = <i> comp (<j> Id B (e.1 x) (e.1 (secEq A B e x @ -i /\ -j)))
+ = <i> comp (<j> Path B (e.1 x) (e.1 (secEq A B e x @ -i /\ -j)))
(((e.2 (e.1 x)).2 (x,<_>e.1 x)@-i).2)
[(i=0)-><j><k>e.1 (secEq A B e x@-k\/-j)
,(i=1)-><j><k>retEq A B e (e.1 x)@-k
where
F (x:Sigma A B):Sigma A' B' = (e.1 x.1, (f x.1).1 x.2)
G (x:Sigma A' B'):Sigma A B = ((e.2 x.1).1.1, ((f (e.2 x.1).1.1).2 (transport (<i> B' ((e.2 x.1).1.2 @ i)) x.2)).1.1)
- FG (x:Sigma A' B'):Id (Sigma A' B') (F (G x)) x
+ FG (x:Sigma A' B'):Path (Sigma A' B') (F (G x)) x
= <i> ((e.2 x.1).1.2 @ -i
,comp (<_> B' ((e.2 x.1).1.2 @ -i))
(transport (<j> B' ((e.2 x.1).1.2 @ j/\-i)) x.2)
[(i=0)-><j>((f (e.2 x.1).1.1).2 (transport (<i> B' ((e.2 x.1).1.2 @ i)) x.2)).1.2 @ j
,(i=1)-><j>transRefl (B' ((e.2 x.1).1.2 @ 0)) x.2@j
])
- GF (x:Sigma A B):Id (Sigma A B) (G (F x)) x
+ GF (x:Sigma A B):Path (Sigma A B) (G (F x)) x
= <i> (secEq A A' e x.1 @ i
,comp (<_> B (secEq A A' e x.1 @ i))
((f (secEq A A' e x.1 @ i)).2 (transport (<j> B' (e.1 (secEq A A' e x.1 @ i\/-j))) ((f x.1).1 x.2))).1.1
[(i=0)-><k>((f (secEq A A' e x.1 @ 0)).2 (transport (<i> B' (lemEquivCoh A A' e x.1@k@i)) ((f x.1).1 x.2))).1.1
- ,(i=1)-><k>compId (B x.1)
+ ,(i=1)-><k>compPath (B x.1)
((f x.1).2 (transport (<j> B' (e.1 x.1)) ((f x.1).1 x.2))).1.1
((f x.1).2 ((f x.1).1 x.2)).1.1
x.2
sigEquivLem' (A A':U) (B:A->U) (B':A'->U)
(e:equiv A A')
- (f:(x:A)->Id U (B x) (B' (e.1 x)))
- : Id U (Sigma A B) (Sigma A' B')
- = transEquivToId (Sigma A B) (Sigma A' B') (sigEquivLem A A' B B' e (\(x:A)->transEquiv' (B' (e.1 x)) (B x) (f x)))
-
-catIso321 (A B : precategory) (e1:Id U (cA A) (cA B)) : U
- = (x y:cA A) -> Id U (cH A x y) (cH B (transport e1 x) (transport e1 y))
-catIso322 (A B : precategory) (e1:Id U (cA A) (cA B)) (e2:catIso321 A B e1) : U
- = (x:cA A) -> Id (cH B (transport e1 x) (transport e1 x))
- (transport (e2 x x) (cId A x))
- (cId B (transport e1 x))
-catIso323 (A B : precategory) (e1:Id U (cA A) (cA B)) (e2:catIso321 A B e1) : U
+ (f:(x:A)->Path U (B x) (B' (e.1 x)))
+ : Path U (Sigma A B) (Sigma A' B')
+ = transEquivToPath (Sigma A B) (Sigma A' B') (sigEquivLem A A' B B' e (\(x:A)->transEquiv' (B' (e.1 x)) (B x) (f x)))
+
+catIso321 (A B : precategory) (e1:Path U (cA A) (cA B)) : U
+ = (x y:cA A) -> Path U (cH A x y) (cH B (transport e1 x) (transport e1 y))
+catIso322 (A B : precategory) (e1:Path U (cA A) (cA B)) (e2:catIso321 A B e1) : U
+ = (x:cA A) -> Path (cH B (transport e1 x) (transport e1 x))
+ (transport (e2 x x) (cPath A x))
+ (cPath B (transport e1 x))
+catIso323 (A B : precategory) (e1:Path U (cA A) (cA B)) (e2:catIso321 A B e1) : U
= (x y z:cA A)(f:cH A x y)(g:cH A y z)->
- Id (cH B (transport e1 x) (transport e1 z))
+ Path (cH B (transport e1 x) (transport e1 z))
(transport (e2 x z) (cC A x y z f g))
(cC B (transport e1 x) (transport e1 y) (transport e1 z)
(transport (e2 x y) f)
(transport (e2 y z) g))
-catIso32 (A B : precategory) : U = (e1 : Id U (cA A) (cA B))
+catIso32 (A B : precategory) : U = (e1 : Path U (cA A) (cA B))
* (e2 : catIso321 A B e1)
* (_ : catIso322 A B e1 e2)
* ( catIso323 A B e1 e2)
-eCatIso32 (A B : precategory) : Id U (catIso31 A B) (catIso32 A B)
- = <i> (e1 : Id U (cA A) (cA B))
+eCatIso32 (A B : precategory) : Path U (catIso31 A B) (catIso32 A B)
+ = <i> (e1 : Path U (cA A) (cA B))
* (let F(e2:catIso311 A B e1):catIso321 A B e1 = \(x y:cA A)-><i>(e2@i) x y in
- transEquivToId
+ transEquivToPath
((e2 : catIso311 A B e1) * (_ : catIso312 A B e1 e2) * ( catIso313 A B e1 e2))
((e2 : catIso321 A B e1) * (_ : catIso322 A B e1 e2) * ( catIso323 A B e1 e2))
(sigEquivLem (catIso311 A B e1) (catIso321 A B e1)
iso (C : precategory) (A B : cA C) : U
= (f : cH C A B) * (g : cH C B A)
- * (_ : Id (cH C A A) (cC C A B A f g) (cId C A))
- * (Id (cH C B B) (cC C B A B g f) (cId C B))
+ * (_ : Path (cH C A A) (cC C A B A f g) (cPath C A))
+ * (Path (cH C B B) (cC C B A B g f) (cPath C B))
-isoEq (C : precategory) (A B : cA C) (f g : iso C A B) (p1 : Id (cH C A B) f.1 g.1) (p2 : Id (cH C B A) f.2.1 g.2.1)
- : Id (iso C A B) f g
+isoEq (C : precategory) (A B : cA C) (f g : iso C A B) (p1 : Path (cH C A B) f.1 g.1) (p2 : Path (cH C B A) f.2.1 g.2.1)
+ : Path (iso C A B) f g
= <i> (p1@i,p2@i
- ,lemIdPProp (Id (cH C A A) (cC C A B A (p1@0) (p2@0)) (cId C A))
- (Id (cH C A A) (cC C A B A (p1@1) (p2@1)) (cId C A))
- (cHSet C A A (cC C A B A (p1@0) (p2@0)) (cId C A))
- (<i>Id (cH C A A) (cC C A B A (p1@i) (p2@i)) (cId C A))
+ ,lemPathPProp (Path (cH C A A) (cC C A B A (p1@0) (p2@0)) (cPath C A))
+ (Path (cH C A A) (cC C A B A (p1@1) (p2@1)) (cPath C A))
+ (cHSet C A A (cC C A B A (p1@0) (p2@0)) (cPath C A))
+ (<i>Path (cH C A A) (cC C A B A (p1@i) (p2@i)) (cPath C A))
f.2.2.1 g.2.2.1 @ i
- ,lemIdPProp (Id (cH C B B) (cC C B A B (p2@0) (p1@0)) (cId C B))
- (Id (cH C B B) (cC C B A B (p2@1) (p1@1)) (cId C B))
- (cHSet C B B (cC C B A B (p2@0) (p1@0)) (cId C B))
- (<i>Id (cH C B B) (cC C B A B (p2@i) (p1@i)) (cId C B))
+ ,lemPathPProp (Path (cH C B B) (cC C B A B (p2@0) (p1@0)) (cPath C B))
+ (Path (cH C B B) (cC C B A B (p2@1) (p1@1)) (cPath C B))
+ (cHSet C B B (cC C B A B (p2@0) (p1@0)) (cPath C B))
+ (<i>Path (cH C B B) (cC C B A B (p2@i) (p1@i)) (cPath C B))
f.2.2.2 g.2.2.2 @ i)
opaque isoEq
invIso (C : precategory) (A B : cA C) (f : iso C A B) : iso C B A = (f.2.1, f.1, f.2.2.2, f.2.2.1)
-idIso (C : precategory) (A : cA C) : iso C A A = (cId C A, cId C A, cIdL C A A (cId C A), cIdL C A A (cId C A))
+idIso (C : precategory) (A : cA C) : iso C A A = (cPath C A, cPath C A, cPathL C A A (cPath C A), cPathL C A A (cPath C A))
invIsoEquiv (C : precategory) (A B : cA C) : isEquiv (iso C A B) (iso C B A) (invIso C A B)
= gradLemma (iso C A B) (iso C B A) (invIso C A B) (invIso C B A) (\(x : iso C B A) -> <_> x) (\(x : iso C A B) -> <_> x)
-invIsoEq (C : precategory) (A B : cA C) : Id U (iso C A B) (iso C B A)
- = equivId (iso C A B) (iso C B A) (invIso C A B) (invIsoEquiv C A B)
+invIsoEq (C : precategory) (A B : cA C) : Path U (iso C A B) (iso C B A)
+ = equivPath (iso C A B) (iso C B A) (invIso C A B) (invIsoEquiv C A B)
compIsoHelper (X : precategory) (A B C : cA X) (f : iso X A B) (g : iso X B C)
- : Id (cH X A A) (cC X A C A (cC X A B C f.1 g.1) (cC X C B A g.2.1 f.2.1)) (cId X A)
- = compId (cH X A A)
+ : Path (cH X A A) (cC X A C A (cC X A B C f.1 g.1) (cC X C B A g.2.1 f.2.1)) (cPath X A)
+ = compPath (cH X A A)
(cC X A C A (cC X A B C f.1 g.1) (cC X C B A g.2.1 f.2.1))
(cC X A B A f.1 (cC X B C A g.1 (cC X C B A g.2.1 f.2.1)))
- (cId X A)
- (cIdC X A B C A f.1 g.1 (cC X C B A g.2.1 f.2.1))
- (compId (cH X A A)
+ (cPath X A)
+ (cPathC X A B C A f.1 g.1 (cC X C B A g.2.1 f.2.1))
+ (compPath (cH X A A)
(cC X A B A f.1 (cC X B C A g.1 (cC X C B A g.2.1 f.2.1)))
(cC X A B A f.1 (cC X B B A (cC X B C B g.1 g.2.1) f.2.1))
- (cId X A)
- (<i>cC X A B A f.1 (cIdC X B C B A g.1 g.2.1 f.2.1 @ -i))
- (compId (cH X A A)
+ (cPath X A)
+ (<i>cC X A B A f.1 (cPathC X B C B A g.1 g.2.1 f.2.1 @ -i))
+ (compPath (cH X A A)
(cC X A B A f.1 (cC X B B A (cC X B C B g.1 g.2.1) f.2.1))
- (cC X A B A f.1 (cC X B B A (cId X B) f.2.1))
- (cId X A)
+ (cC X A B A f.1 (cC X B B A (cPath X B) f.2.1))
+ (cPath X A)
(<i>cC X A B A f.1 (cC X B B A (g.2.2.1 @ i) f.2.1))
- (compId (cH X A A)
- (cC X A B A f.1 (cC X B B A (cId X B) f.2.1))
+ (compPath (cH X A A)
+ (cC X A B A f.1 (cC X B B A (cPath X B) f.2.1))
(cC X A B A f.1 f.2.1)
- (cId X A)
- (<i>cC X A B A f.1 (cIdL X B A f.2.1 @ i))
+ (cPath X A)
+ (<i>cC X A B A f.1 (cPathL X B A f.2.1 @ i))
f.2.2.1)))
opaque compIsoHelper
,compIsoHelper X A B C f g
,compIsoHelper X C B A (invIso X B C g) (invIso X A B f))
-IdLIso (C : precategory) (A B : cA C) (f : iso C A B) : Id (iso C A B) (compIso C A A B (idIso C A) f) f
- = isoEq C A B (compIso C A A B (idIso C A) f) f (cIdL C A B f.1) (cIdR C B A f.2.1)
-opaque IdLIso
-IdRIso (C : precategory) (A B : cA C) (f : iso C A B) : Id (iso C A B) (compIso C A B B f (idIso C B)) f
- = isoEq C A B (compIso C A B B f (idIso C B)) f (cIdR C A B f.1) (cIdL C B A f.2.1)
-opaque IdRIso
-IdCIso (X : precategory) (A B C D : cA X) (f : iso X A B) (g : iso X B C) (h : iso X C D)
- : Id (iso X A D) (compIso X A C D (compIso X A B C f g) h) (compIso X A B D f (compIso X B C D g h))
+PathLIso (C : precategory) (A B : cA C) (f : iso C A B) : Path (iso C A B) (compIso C A A B (idIso C A) f) f
+ = isoEq C A B (compIso C A A B (idIso C A) f) f (cPathL C A B f.1) (cPathR C B A f.2.1)
+opaque PathLIso
+PathRIso (C : precategory) (A B : cA C) (f : iso C A B) : Path (iso C A B) (compIso C A B B f (idIso C B)) f
+ = isoEq C A B (compIso C A B B f (idIso C B)) f (cPathR C A B f.1) (cPathL C B A f.2.1)
+opaque PathRIso
+PathCIso (X : precategory) (A B C D : cA X) (f : iso X A B) (g : iso X B C) (h : iso X C D)
+ : Path (iso X A D) (compIso X A C D (compIso X A B C f g) h) (compIso X A B D f (compIso X B C D g h))
= isoEq X A D (compIso X A C D (compIso X A B C f g) h) (compIso X A B D f (compIso X B C D g h))
- (cIdC X A B C D f.1 g.1 h.1) (<i>cIdC X D C B A h.2.1 g.2.1 f.2.1@-i)
-opaque IdCIso
-IdInvLIso (X : precategory) (A B : cA X) (f : iso X A B) : Id (iso X B B) (compIso X B A B (invIso X A B f) f) (idIso X B)
+ (cPathC X A B C D f.1 g.1 h.1) (<i>cPathC X D C B A h.2.1 g.2.1 f.2.1@-i)
+opaque PathCIso
+PathInvLIso (X : precategory) (A B : cA X) (f : iso X A B) : Path (iso X B B) (compIso X B A B (invIso X A B f) f) (idIso X B)
= isoEq X B B (compIso X B A B (invIso X A B f) f) (idIso X B) f.2.2.2 f.2.2.2
-opaque IdInvLIso
+opaque PathInvLIso
opaque compIso
-eqToIso (C : precategory) (A B : cA C) (p : Id (cA C) A B) : iso C A B
- = J (cA C) A (\(B : cA C) -> \(p : Id (cA C) A B) -> iso C A B) (idIso C A) B p
+eqToIso (C : precategory) (A B : cA C) (p : Path (cA C) A B) : iso C A B
+ = J (cA C) A (\(B : cA C) -> \(p : Path (cA C) A B) -> iso C A B) (idIso C A) B p
isCategory (C : precategory) : U = (A : cA C) -> isContr ((B : cA C) * iso C A B)
propIsCategory (C : precategory) : prop (isCategory C)
(\(A : cA C) -> propIsContr ((B : cA C) * iso C A B))
category : U = (C:precategory)*isCategory C
-lemIsCategory (C : precategory) (isC : isCategory C) (A B : cA C) (e : iso C A B) : Id ((B : cA C) * iso C A B) (A, idIso C A) (B, e)
+lemIsCategory (C : precategory) (isC : isCategory C) (A B : cA C) (e : iso C A B) : Path ((B : cA C) * iso C A B) (A, idIso C A) (B, e)
= <i> (isContrProp ((B : cA C) * iso C A B) (isC A) (A, idIso C A) (B, e) @ i)
-lemIsCategory2 (C : precategory) (isC : isCategory C) (A B : cA C) : isEquiv (Id (cA C) A B) (iso C A B) (eqToIso C A B)
- = equivFunFib (cA C) (Id (cA C) A) (iso C A) (eqToIso C A) (lemSinglContr (cA C) A) (isC A) B
+lemIsCategory2 (C : precategory) (isC : isCategory C) (A B : cA C) : isEquiv (Path (cA C) A B) (iso C A B) (eqToIso C A B)
+ = equivFunFib (cA C) (Path (cA C) A) (iso C A) (eqToIso C A) (lemSinglContr (cA C) A) (isC A) B
-lemIsCategory3 (C : precategory) (isC : isCategory C) (A B : cA C) : Id U (Id (cA C) A B) (iso C A B)
- = equivId (Id (cA C) A B) (iso C A B) (eqToIso C A B) (lemIsCategory2 C isC A B)
+lemIsCategory3 (C : precategory) (isC : isCategory C) (A B : cA C) : Path U (Path (cA C) A B) (iso C A B)
+ = equivPath (Path (cA C) A B) (iso C A B) (eqToIso C A B) (lemIsCategory2 C isC A B)
--
= (P : cA X -> U)
* (H : (x y : cA X) -> (a : P x) -> (b : P y) -> (f : cH X x y) -> U)
* (propH : (x y : cA X) -> (a : P x) -> (b : P y) -> (f : cH X x y) -> prop (H x y a b f))
- * (Hid : (x : cA X) -> (a : P x) -> H x x a a (cId X x))
+ * (Hid : (x : cA X) -> (a : P x) -> H x x a a (cPath X x))
* ((x y z : cA X) -> (a : P x) -> (b : P y) -> (c : P z) ->
(f : cH X x y) -> (g : cH X y z) ->
H x y a b f -> H y z b c g -> H x z a c (cC X x y z f g))
sH (X : precategory) (S : structure X) : (x y : cA X) -> (a : sP X S x) -> (b : sP X S y) -> (f : cH X x y) -> U = S.2.1
sHProp (X : precategory) (S : structure X)
: (x y : cA X) -> (a : sP X S x) -> (b : sP X S y) -> (f : cH X x y) -> prop (sH X S x y a b f) = S.2.2.1
-sHId (X : precategory) (S : structure X) : (x : cA X) -> (a : sP X S x) -> sH X S x x a a (cId X x) = S.2.2.2.1
+sHPath (X : precategory) (S : structure X) : (x : cA X) -> (a : sP X S x) -> sH X S x x a a (cPath X x) = S.2.2.2.1
sHComp (X : precategory) (S : structure X)
: (x y z : cA X) -> (a : sP X S x) -> (b : sP X S y) -> (c : sP X S z) ->
(f : cH X x y) -> (g : cH X y z) ->
isStandardStructure (X : precategory) (S : structure X) : U
= (x : cA X) -> (a b : sP X S x) ->
- sH X S x x a b (cId X x) -> sH X S x x b a (cId X x) ->
- Id (sP X S x) a b
+ sH X S x x a b (cPath X x) -> sH X S x x b a (cPath X x) ->
+ Path (sP X S x) a b
sipPrecategory (C : precategory) (S : structure C) : precategory = ((ob, hom), (id, cmp, homSet, idL, idR, idC))
where
homSet (X Y : ob) : set (hom X Y)
= setSig (cH C X.1 Y.1) (sH C S X.1 Y.1 X.2 Y.2)
(cHSet C X.1 Y.1) (\(f : cH C X.1 Y.1) -> propSet (sH C S X.1 Y.1 X.2 Y.2 f) (sHProp C S X.1 Y.1 X.2 Y.2 f))
- id (X : ob) : hom X X = (cId C X.1, sHId C S X.1 X.2)
+ id (X : ob) : hom X X = (cPath C X.1, sHPath C S X.1 X.2)
cmp (x y z : ob) (f : hom x y) (g : hom y z) : hom x z
= (cC C x.1 y.1 z.1 f.1 g.1, sHComp C S x.1 y.1 z.1 x.2 y.2 z.2 f.1 g.1 f.2 g.2)
- idL (x y : ob) (f : hom x y) : Id (hom x y) (cmp x x y (id x) f) f
- = <i> (cIdL C x.1 y.1 f.1 @ i
- ,lemIdPProp (sH C S x.1 y.1 x.2 y.2 (cmp x x y (id x) f).1)
+ idL (x y : ob) (f : hom x y) : Path (hom x y) (cmp x x y (id x) f) f
+ = <i> (cPathL C x.1 y.1 f.1 @ i
+ ,lemPathPProp (sH C S x.1 y.1 x.2 y.2 (cmp x x y (id x) f).1)
(sH C S x.1 y.1 x.2 y.2 f.1)
(sHProp C S x.1 y.1 x.2 y.2 (cmp x x y (id x) f).1)
- (<i>sH C S x.1 y.1 x.2 y.2 (cIdL C x.1 y.1 f.1 @ i))
+ (<i>sH C S x.1 y.1 x.2 y.2 (cPathL C x.1 y.1 f.1 @ i))
(cmp x x y (id x) f).2 f.2 @ i)
- idR (x y : ob) (f : hom x y) : Id (hom x y) (cmp x y y f (id y)) f
- = <i> (cIdR C x.1 y.1 f.1 @ i
- ,lemIdPProp (sH C S x.1 y.1 x.2 y.2 (cmp x y y f (id y)).1)
+ idR (x y : ob) (f : hom x y) : Path (hom x y) (cmp x y y f (id y)) f
+ = <i> (cPathR C x.1 y.1 f.1 @ i
+ ,lemPathPProp (sH C S x.1 y.1 x.2 y.2 (cmp x y y f (id y)).1)
(sH C S x.1 y.1 x.2 y.2 f.1)
(sHProp C S x.1 y.1 x.2 y.2 (cmp x y y f (id y)).1)
- (<i>sH C S x.1 y.1 x.2 y.2 (cIdR C x.1 y.1 f.1 @ i))
+ (<i>sH C S x.1 y.1 x.2 y.2 (cPathR C x.1 y.1 f.1 @ i))
(cmp x y y f (id y)).2 f.2 @ i)
- idC (x y z w : ob) (f : hom x y) (g : hom y z) (h : hom z w) : Id (hom x w) (cmp x z w (cmp x y z f g) h) (cmp x y w f (cmp y z w g h))
- = <i> (cIdC C x.1 y.1 z.1 w.1 f.1 g.1 h.1 @ i
- ,lemIdPProp (sH C S x.1 w.1 x.2 w.2 (cmp x z w (cmp x y z f g) h).1)
+ idC (x y z w : ob) (f : hom x y) (g : hom y z) (h : hom z w) : Path (hom x w) (cmp x z w (cmp x y z f g) h) (cmp x y w f (cmp y z w g h))
+ = <i> (cPathC C x.1 y.1 z.1 w.1 f.1 g.1 h.1 @ i
+ ,lemPathPProp (sH C S x.1 w.1 x.2 w.2 (cmp x z w (cmp x y z f g) h).1)
(sH C S x.1 w.1 x.2 w.2 (cmp x y w f (cmp y z w g h)).1)
(sHProp C S x.1 w.1 x.2 w.2 (cmp x z w (cmp x y z f g) h).1)
- (<i>sH C S x.1 w.1 x.2 w.2 (cIdC C x.1y.1 z.1 w.1 f.1 g.1 h.1 @ i))
+ (<i>sH C S x.1 w.1 x.2 w.2 (cPathC C x.1y.1 z.1 w.1 f.1 g.1 h.1 @ i))
(cmp x z w (cmp x y z f g) h).2 (cmp x y w f (cmp y z w g h)).2 @ i)
-isContrProp (A : U) (p : isContr A) (x y : A) : Id A x y = compId A x p.1 y (<i> p.2 x @ -i) (p.2 y)
+isContrProp (A : U) (p : isContr A) (x y : A) : Path A x y = compPath A x p.1 y (<i> p.2 x @ -i) (p.2 y)
isContrSig (A:U) (B:A-> U) (cA:isContr A) (cB : (x:A) -> isContr (B x)) : isContr (Sigma A B)
= ((cA.1, (cB cA.1).1), \(x:Sigma A B) -> propSig A B (isContrProp A cA) (\(x:A)->isContrProp (B x) (cB x)) (cA.1, (cB cA.1).1) x)
-- where
-- D : precategory = sipPrecategory X S
-- eq1 (A : cA D)
- -- : Id U ((B : cA D) * iso D A B)
+ -- : Path U ((B : cA D) * iso D A B)
-- ((C : (B1 : cA X) * ( iso X A.1 B1))
-- * (B2 : sP X S C.1) * (_ : sH X S A.1 C.1 A.2 B2 C.2.1) * sH X S C.1 A.1 B2 A.2 C.2.2.1)
- -- = isoId
+ -- = isoPath
-- ((B : cA D) * iso D A B)
-- ((C : (B1 : cA X) * ( iso X A.1 B1))
-- * (B2 : sP X S C.1) * (_ : sH X S A.1 C.1 A.2 B2 C.2.1) * sH X S C.1 A.1 B2 A.2 C.2.2.1)
-- : (B : cA D) * iso D A B
-- = ((C.1.1, C.2.1), (C.1.2.1, C.2.2.1), (C.1.2.2.1, C.2.2.2)
-- , <i> (C.1.2.2.2.1 @ i
- -- ,lemIdPProp (sH X S A.1 A.1 A.2 A.2 (cC X A.1 C.1.1 A.1 C.1.2.1 C.1.2.2.1))
- -- (sH X S A.1 A.1 A.2 A.2 (cId X A.1))
+ -- ,lemPathPProp (sH X S A.1 A.1 A.2 A.2 (cC X A.1 C.1.1 A.1 C.1.2.1 C.1.2.2.1))
+ -- (sH X S A.1 A.1 A.2 A.2 (cPath X A.1))
-- (sHProp X S A.1 A.1 A.2 A.2 (cC X A.1 C.1.1 A.1 C.1.2.1 C.1.2.2.1))
-- (<i>sH X S A.1 A.1 A.2 A.2 (C.1.2.2.2.1 @ i))
-- (sHComp X S A.1 C.1.1 A.1 A.2 C.2.1 A.2 C.1.2.1 C.1.2.2.1 C.2.2.1 C.2.2.2)
- -- (sHId X S A.1 A.2) @ i)
+ -- (sHPath X S A.1 A.2) @ i)
-- , <i> (C.1.2.2.2.2 @ i
- -- ,lemIdPProp (sH X S C.1.1 C.1.1 C.2.1 C.2.1 (cC X C.1.1 A.1 C.1.1 C.1.2.2.1 C.1.2.1))
- -- (sH X S C.1.1 C.1.1 C.2.1 C.2.1 (cId X C.1.1))
+ -- ,lemPathPProp (sH X S C.1.1 C.1.1 C.2.1 C.2.1 (cC X C.1.1 A.1 C.1.1 C.1.2.2.1 C.1.2.1))
+ -- (sH X S C.1.1 C.1.1 C.2.1 C.2.1 (cPath X C.1.1))
-- (sHProp X S C.1.1 C.1.1 C.2.1 C.2.1 (cC X C.1.1 A.1 C.1.1 C.1.2.2.1 C.1.2.1))
-- (<i>sH X S C.1.1 C.1.1 C.2.1 C.2.1 (C.1.2.2.2.2 @ i))
-- (sHComp X S C.1.1 A.1 C.1.1 C.2.1 A.2 C.2.1 C.1.2.2.1 C.1.2.1 C.2.2.2 C.2.2.1)
- -- (sHId X S C.1.1 C.2.1) @ i))
+ -- (sHPath X S C.1.1 C.2.1) @ i))
-- FG (C : ((C : (B1 : cA X) * ( iso X A.1 B1))
-- * (B2 : sP X S C.1) * (_ : sH X S A.1 C.1 A.2 B2 C.2.1) * sH X S C.1 A.1 B2 A.2 C.2.2.1))
- -- : Id ((C : (B1 : cA X) * ( iso X A.1 B1))
+ -- : Path ((C : (B1 : cA X) * ( iso X A.1 B1))
-- * (B2 : sP X S C.1) * (_ : sH X S A.1 C.1 A.2 B2 C.2.1) * sH X S C.1 A.1 B2 A.2 C.2.2.1)
-- (F (G C)) C
-- = <_> C
- -- GF (C : (B : cA D) * iso D A B) : Id ((B : cA D) * iso D A B) (G (F C)) C
+ -- GF (C : (B : cA D) * iso D A B) : Path ((B : cA D) * iso D A B) (G (F C)) C
-- = <i> (C.1, C.2.1, C.2.2.1
-- , <j> ((C.2.2.2.1 @ j).1
- -- ,lemIdPSet (sH X S A.1 A.1 A.2 A.2 (cC X A.1 C.1.1 A.1 C.2.1.1 C.2.2.1.1))
- -- (sH X S A.1 A.1 A.2 A.2 (cId X A.1))
+ -- ,lemPathPSet (sH X S A.1 A.1 A.2 A.2 (cC X A.1 C.1.1 A.1 C.2.1.1 C.2.2.1.1))
+ -- (sH X S A.1 A.1 A.2 A.2 (cPath X A.1))
-- (propSet (sH X S A.1 A.1 A.2 A.2 (cC X A.1 C.1.1 A.1 C.2.1.1 C.2.2.1.1))
-- (sHProp X S A.1 A.1 A.2 A.2 (cC X A.1 C.1.1 A.1 C.2.1.1 C.2.2.1.1)))
-- (<j>sH X S A.1 A.1 A.2 A.2 (C.2.2.2.1 @ j).1)
-- (sHComp X S A.1 C.1.1 A.1 A.2 C.1.2 A.2 C.2.1.1 C.2.2.1.1 C.2.1.2 C.2.2.1.2)
- -- (sHId X S A.1 A.2)
+ -- (sHPath X S A.1 A.2)
-- (<j>((G (F C)).2.2.2.1 @ j).2)
-- (<j>(C.2.2.2.1 @ j).2) @i@j)
-- , <j> ((C.2.2.2.2 @ j).1
- -- ,lemIdPSet (sH X S C.1.1 C.1.1 C.1.2 C.1.2 (cC X C.1.1 A.1 C.1.1 C.2.2.1.1 C.2.1.1))
- -- (sH X S C.1.1 C.1.1 C.1.2 C.1.2 (cId X C.1.1))
+ -- ,lemPathPSet (sH X S C.1.1 C.1.1 C.1.2 C.1.2 (cC X C.1.1 A.1 C.1.1 C.2.2.1.1 C.2.1.1))
+ -- (sH X S C.1.1 C.1.1 C.1.2 C.1.2 (cPath X C.1.1))
-- (propSet (sH X S C.1.1 C.1.1 C.1.2 C.1.2 (cC X C.1.1 A.1 C.1.1 C.2.2.1.1 C.2.1.1))
-- (sHProp X S C.1.1 C.1.1 C.1.2 C.1.2 (cC X C.1.1 A.1 C.1.1 C.2.2.1.1 C.2.1.1)))
-- (<j>sH X S C.1.1 C.1.1 C.1.2 C.1.2 (C.2.2.2.2 @ j).1)
-- (sHComp X S C.1.1 A.1 C.1.1 C.1.2 A.2 C.1.2 C.2.2.1.1 C.2.1.1 C.2.2.1.2 C.2.1.2)
- -- (sHId X S C.1.1 C.1.2)
+ -- (sHPath X S C.1.1 C.1.2)
-- (<j>((G (F C)).2.2.2.2 @ j).2)
-- (<j>(C.2.2.2.2 @ j).2) @i@j))
-- hole0 (A : cA D)
-- (B2 : sP X S C.1) * (_ : sH X S A.1 C.1 A.2 B2 C.2.1) * sH X S C.1 A.1 B2 A.2 C.2.2.1)
-- (isC A.1)
-- (\(C : (B1 : cA X) * ( iso X A.1 B1)) ->
- -- let p : Id ((B1 : cA X) * ( iso X A.1 B1)) (A.1, idIso X A.1) C
+ -- let p : Path ((B1 : cA X) * ( iso X A.1 B1)) (A.1, idIso X A.1) C
-- = lemIsCategory X isC A.1 C.1 C.2
-- in transport
-- (<i> isContr ((B2 : sP X S (p@i).1) * (_ : sH X S A.1 (p@i).1 A.2 B2 (p@i).2.1) * sH X S (p@i).1 A.1 B2 A.2 (p@i).2.2.1))
- -- ((A.2,sHId X S A.1 A.2,sHId X S A.1 A.2)
- -- ,\(Y : (B2 : sP X S A.1) * (_ : sH X S A.1 A.1 A.2 B2 (cId X A.1)) * sH X S A.1 A.1 B2 A.2 (cId X A.1)) ->
+ -- ((A.2,sHPath X S A.1 A.2,sHPath X S A.1 A.2)
+ -- ,\(Y : (B2 : sP X S A.1) * (_ : sH X S A.1 A.1 A.2 B2 (cPath X A.1)) * sH X S A.1 A.1 B2 A.2 (cPath X A.1)) ->
-- <i> (isS A.1 A.2 Y.1 Y.2.1 Y.2.2 @ i
- -- ,lemIdPProp (sH X S A.1 A.1 A.2 A.2 (cId X A.1))
- -- (sH X S A.1 A.1 A.2 Y.1 (cId X A.1))
- -- (sHProp X S A.1 A.1 A.2 A.2 (cId X A.1))
- -- (<i>sH X S A.1 A.1 A.2 (isS A.1 A.2 Y.1 Y.2.1 Y.2.2 @ i) (cId X A.1))
- -- (sHId X S A.1 A.2) Y.2.1 @ i
- -- ,lemIdPProp (sH X S A.1 A.1 A.2 A.2 (cId X A.1))
- -- (sH X S A.1 A.1 Y.1 A.2 (cId X A.1))
- -- (sHProp X S A.1 A.1 A.2 A.2 (cId X A.1))
- -- (<i>sH X S A.1 A.1 (isS A.1 A.2 Y.1 Y.2.1 Y.2.2 @ i) A.2 (cId X A.1))
- -- (sHId X S A.1 A.2) Y.2.2 @ i)))
+ -- ,lemPathPProp (sH X S A.1 A.1 A.2 A.2 (cPath X A.1))
+ -- (sH X S A.1 A.1 A.2 Y.1 (cPath X A.1))
+ -- (sHProp X S A.1 A.1 A.2 A.2 (cPath X A.1))
+ -- (<i>sH X S A.1 A.1 A.2 (isS A.1 A.2 Y.1 Y.2.1 Y.2.2 @ i) (cPath X A.1))
+ -- (sHPath X S A.1 A.2) Y.2.1 @ i
+ -- ,lemPathPProp (sH X S A.1 A.1 A.2 A.2 (cPath X A.1))
+ -- (sH X S A.1 A.1 Y.1 A.2 (cPath X A.1))
+ -- (sHProp X S A.1 A.1 A.2 A.2 (cPath X A.1))
+ -- (<i>sH X S A.1 A.1 (isS A.1 A.2 Y.1 Y.2.1 Y.2.2 @ i) A.2 (cPath X A.1))
+ -- (sHPath X S A.1 A.2) Y.2.2 @ i)))
-- hole (A : cA D) : isContr ((B : cA D) * iso D A B) = transport (<i>isContr(eq1 A@-i)) (hole0 A)
opaque sip
functor (A B : precategory) : U
= (Fob : cA A -> cA B)
* (Fmor : (x y : cA A) -> cH A x y -> cH B (Fob x) (Fob y))
- * (Fid : (x : cA A) -> Id (cH B (Fob x) (Fob x)) (Fmor x x (cId A x)) (cId B (Fob x)))
- * ((x y z : cA A) -> (f : cH A x y) -> (g : cH A y z) -> Id (cH B (Fob x) (Fob z)) (Fmor x z (cC A x y z f g)) (cC B (Fob x) (Fob y) (Fob z) (Fmor x y f) (Fmor y z g)))
+ * (Fid : (x : cA A) -> Path (cH B (Fob x) (Fob x)) (Fmor x x (cPath A x)) (cPath B (Fob x)))
+ * ((x y z : cA A) -> (f : cH A x y) -> (g : cH A y z) -> Path (cH B (Fob x) (Fob z)) (Fmor x z (cC A x y z f g)) (cC B (Fob x) (Fob y) (Fob z) (Fmor x y f) (Fmor y z g)))
idFunctor (A : precategory) : functor A A
= (\(x : cA A) -> x
,\(x y : cA A) (h : cH A x y) -> h
- ,\(x : cA A) -> <_> cId A x
+ ,\(x : cA A) -> <_> cPath A x
,\(x y z : cA A) (f : cH A x y) (g : cH A y z) -> <_> cC A x y z f g)
compFunctor (A B C : precategory) (F : functor A B) (G : functor B C) : functor A C
= (\(x : cA A) -> G.1 (F.1 x)
,\(x y : cA A) (h : cH A x y) -> G.2.1 (F.1 x) (F.1 y) (F.2.1 x y h)
- ,\(x : cA A) -> compId (cH C (G.1 (F.1 x)) (G.1 (F.1 x)))
- (G.2.1 (F.1 x) (F.1 x) (F.2.1 x x (cId A x)))
- (G.2.1 (F.1 x) (F.1 x) (cId B (F.1 x)))
- (cId C (G.1 (F.1 x)))
+ ,\(x : cA A) -> compPath (cH C (G.1 (F.1 x)) (G.1 (F.1 x)))
+ (G.2.1 (F.1 x) (F.1 x) (F.2.1 x x (cPath A x)))
+ (G.2.1 (F.1 x) (F.1 x) (cPath B (F.1 x)))
+ (cPath C (G.1 (F.1 x)))
(<i>G.2.1 (F.1 x) (F.1 x) (F.2.2.1 x @ i))
(G.2.2.1 (F.1 x))
,\(x y z : cA A) (f : cH A x y) (g : cH A y z) ->
- compId (cH C (G.1 (F.1 x)) (G.1 (F.1 z)))
+ compPath (cH C (G.1 (F.1 x)) (G.1 (F.1 z)))
(G.2.1 (F.1 x) (F.1 z) (F.2.1 x z (cC A x y z f g)))
(G.2.1 (F.1 x) (F.1 z) (cC B (F.1 x) (F.1 y) (F.1 z) (F.2.1 x y f) (F.2.1 y z g)))
(cC C (G.1 (F.1 x)) (G.1 (F.1 y)) (G.1 (F.1 z)) (G.2.1 (F.1 x) (F.1 y) (F.2.1 x y f)) (G.2.1 (F.1 y) (F.1 z) (F.2.1 y z g)))
ffFunctor (A B : precategory) (F : functor A B) : U
= (a b : cA A) -> isEquiv (cH A a b) (cH B (F.1 a) (F.1 b)) (F.2.1 a b)
ffEq (A B : precategory) (F : functor A B) (ff : ffFunctor A B F) (a b : cA A)
- : Id U (cH A a b) (cH B (F.1 a) (F.1 b))
- = equivId (cH A a b) (cH B (F.1 a) (F.1 b)) (F.2.1 a b) (ff a b)
+ : Path U (cH A a b) (cH B (F.1 a) (F.1 b))
+ = equivPath (cH A a b) (cH B (F.1 a) (F.1 b)) (F.2.1 a b) (ff a b)
propFFFunctor (A B : precategory) (F : functor A B) : prop (ffFunctor A B F)
= propPi (cA A) (\(a : cA A) -> (b : cA A) -> isEquiv (cH A a b) (cH B (F.1 a) (F.1 b)) (F.2.1 a b))
(\(a : cA A) -> propPi (cA A) (\(b : cA A) -> isEquiv (cH A a b) (cH B (F.1 a) (F.1 b)) (F.2.1 a b))
(\(b : cA A) -> propIsEquiv (cH A a b) (cH B (F.1 a) (F.1 b)) (F.2.1 a b)))
-lem10 (A B : U) (e : equiv A B) (x y : B) (p : Id A (e.2 x).1.1 (e.2 y).1.1) : Id B x y
+lem10 (A B : U) (e : equiv A B) (x y : B) (p : Path A (e.2 x).1.1 (e.2 y).1.1) : Path B x y
= transport
- (compId U (Id B (e.1 (e.2 x).1.1) (e.1 (e.2 y).1.1)) (Id B x (e.1 (e.2 y).1.1)) (Id B x y)
- (<i> Id B (retEq A B e x @ i) (e.1 (e.2 y).1.1)) (<i> Id B x (retEq A B e y @ i)))
+ (compPath U (Path B (e.1 (e.2 x).1.1) (e.1 (e.2 y).1.1)) (Path B x (e.1 (e.2 y).1.1)) (Path B x y)
+ (<i> Path B (retEq A B e x @ i) (e.1 (e.2 y).1.1)) (<i> Path B x (retEq A B e y @ i)))
(mapOnPath A B e.1 (e.2 x).1.1 (e.2 y).1.1 p)
opaque lem10
-lem10' (A B : U) (e : equiv A B) (x y : A) (p : Id B (e.1 x) (e.1 y)) : Id A x y
+lem10' (A B : U) (e : equiv A B) (x y : A) (p : Path B (e.1 x) (e.1 y)) : Path A x y
= transport
- (compId U (Id A (e.2 (e.1 x)).1.1 (e.2 (e.1 y)).1.1) (Id A x (e.2 (e.1 y)).1.1) (Id A x y)
- (<i> Id A (secEq A B e x @ i) (e.2 (e.1 y)).1.1) (<i> Id A x (secEq A B e y @ i))
+ (compPath U (Path A (e.2 (e.1 x)).1.1 (e.2 (e.1 y)).1.1) (Path A x (e.2 (e.1 y)).1.1) (Path A x y)
+ (<i> Path A (secEq A B e x @ i) (e.2 (e.1 y)).1.1) (<i> Path A x (secEq A B e y @ i))
)
(mapOnPath B A (invEq A B e) (e.1 x) (e.1 y) p)
opaque lem10'
lemFF (A B : precategory) (F : functor A B) (ff : ffFunctor A B F) (x y : cA A)
- : Id U (iso A x y) (iso B (F.1 x) (F.1 y))
+ : Path U (iso A x y) (iso B (F.1 x) (F.1 y))
= hole
where
F0 (f : iso A x y) : iso B (F.1 x) (F.1 y)
= (F.2.1 x y f.1, F.2.1 y x f.2.1
- ,compId (cH B (F.1 x) (F.1 x))
+ ,compPath (cH B (F.1 x) (F.1 x))
(cC B (F.1 x) (F.1 y) (F.1 x) (F.2.1 x y f.1) (F.2.1 y x f.2.1))
(F.2.1 x x (cC A x y x f.1 f.2.1))
- (cId B (F.1 x))
+ (cPath B (F.1 x))
(<i>F.2.2.2 x y x f.1 f.2.1 @-i)
- (compId (cH B (F.1 x) (F.1 x))
+ (compPath (cH B (F.1 x) (F.1 x))
(F.2.1 x x (cC A x y x f.1 f.2.1))
- (F.2.1 x x (cId A x))
- (cId B (F.1 x))
+ (F.2.1 x x (cPath A x))
+ (cPath B (F.1 x))
(<i>F.2.1 x x (f.2.2.1@i))
(F.2.2.1 x))
- ,compId (cH B (F.1 y) (F.1 y))
+ ,compPath (cH B (F.1 y) (F.1 y))
(cC B (F.1 y) (F.1 x) (F.1 y) (F.2.1 y x f.2.1) (F.2.1 x y f.1))
(F.2.1 y y (cC A y x y f.2.1 f.1))
- (cId B (F.1 y))
+ (cPath B (F.1 y))
(<i>F.2.2.2 y x y f.2.1 f.1 @-i)
- (compId (cH B (F.1 y) (F.1 y))
+ (compPath (cH B (F.1 y) (F.1 y))
(F.2.1 y y (cC A y x y f.2.1 f.1))
- (F.2.1 y y (cId A y))
- (cId B (F.1 y))
+ (F.2.1 y y (cPath A y))
+ (cPath B (F.1 y))
(<i>F.2.1 y y (f.2.2.2@i))
(F.2.2.1 y)))
G0 (g : iso B (F.1 x) (F.1 y)) : iso A x y
= ((ff x y g.1).1.1
,(ff y x g.2.1).1.1
- ,lem10' (cH A x x) (cH B (F.1 x) (F.1 x)) (F.2.1 x x, ff x x) (cC A x y x (ff x y g.1).1.1 (ff y x g.2.1).1.1) (cId A x)
- (compId (cH B (F.1 x) (F.1 x))
+ ,lem10' (cH A x x) (cH B (F.1 x) (F.1 x)) (F.2.1 x x, ff x x) (cC A x y x (ff x y g.1).1.1 (ff y x g.2.1).1.1) (cPath A x)
+ (compPath (cH B (F.1 x) (F.1 x))
(F.2.1 x x (cC A x y x (ff x y g.1).1.1 (ff y x g.2.1).1.1))
(cC B (F.1 x) (F.1 y) (F.1 x) (F.2.1 x y (ff x y g.1).1.1) (F.2.1 y x (ff y x g.2.1).1.1))
- (F.2.1 x x (cId A x))
+ (F.2.1 x x (cPath A x))
(F.2.2.2 x y x (ff x y g.1).1.1 (ff y x g.2.1).1.1)
- (compId (cH B (F.1 x) (F.1 x))
+ (compPath (cH B (F.1 x) (F.1 x))
(cC B (F.1 x) (F.1 y) (F.1 x) (F.2.1 x y (ff x y g.1).1.1) (F.2.1 y x (ff y x g.2.1).1.1))
(cC B (F.1 x) (F.1 y) (F.1 x) g.1 g.2.1)
- (F.2.1 x x (cId A x))
+ (F.2.1 x x (cPath A x))
(<i>cC B (F.1 x) (F.1 y) (F.1 x) ((ff x y g.1).1.2@-i) ((ff y x g.2.1).1.2@-i))
- (compId (cH B (F.1 x) (F.1 x))
+ (compPath (cH B (F.1 x) (F.1 x))
(cC B (F.1 x) (F.1 y) (F.1 x) g.1 g.2.1)
- (cId B (F.1 x))
- (F.2.1 x x (cId A x))
+ (cPath B (F.1 x))
+ (F.2.1 x x (cPath A x))
g.2.2.1
(<i>F.2.2.1 x @ -i))))
- ,lem10' (cH A y y) (cH B (F.1 y) (F.1 y)) (F.2.1 y y, ff y y) (cC A y x y (ff y x g.2.1).1.1 (ff x y g.1).1.1) (cId A y)
- (compId (cH B (F.1 y) (F.1 y))
+ ,lem10' (cH A y y) (cH B (F.1 y) (F.1 y)) (F.2.1 y y, ff y y) (cC A y x y (ff y x g.2.1).1.1 (ff x y g.1).1.1) (cPath A y)
+ (compPath (cH B (F.1 y) (F.1 y))
(F.2.1 y y (cC A y x y (ff y x g.2.1).1.1 (ff x y g.1).1.1))
(cC B (F.1 y) (F.1 x) (F.1 y) (F.2.1 y x (ff y x g.2.1).1.1) (F.2.1 x y (ff x y g.1).1.1))
- (F.2.1 y y (cId A y))
+ (F.2.1 y y (cPath A y))
(F.2.2.2 y x y (ff y x g.2.1).1.1 (ff x y g.1).1.1)
- (compId (cH B (F.1 y) (F.1 y))
+ (compPath (cH B (F.1 y) (F.1 y))
(cC B (F.1 y) (F.1 x) (F.1 y) (F.2.1 y x (ff y x g.2.1).1.1) (F.2.1 x y (ff x y g.1).1.1))
(cC B (F.1 y) (F.1 x) (F.1 y) g.2.1 g.1)
- (F.2.1 y y (cId A y))
+ (F.2.1 y y (cPath A y))
(<i>cC B (F.1 y) (F.1 x) (F.1 y) ((ff y x g.2.1).1.2@-i) ((ff x y g.1).1.2@-i))
- (compId (cH B (F.1 y) (F.1 y))
+ (compPath (cH B (F.1 y) (F.1 y))
(cC B (F.1 y) (F.1 x) (F.1 y) g.2.1 g.1)
- (cId B (F.1 y))
- (F.2.1 y y (cId A y))
+ (cPath B (F.1 y))
+ (F.2.1 y y (cPath A y))
g.2.2.2
(<i>F.2.2.1 y @ -i))))
)
- FG (g : iso B (F.1 x) (F.1 y)) : Id (iso B (F.1 x) (F.1 y)) (F0 (G0 g)) g
+ FG (g : iso B (F.1 x) (F.1 y)) : Path (iso B (F.1 x) (F.1 y)) (F0 (G0 g)) g
= isoEq B (F.1 x) (F.1 y) (F0 (G0 g)) g
(<i>(ff x y g.1).1.2@-i) (<i>(ff y x g.2.1).1.2@-i)
- GF (f : iso A x y) : Id (iso A x y) (G0 (F0 f)) f
+ GF (f : iso A x y) : Path (iso A x y) (G0 (F0 f)) f
= isoEq A x y (G0 (F0 f)) f
(<i> ((ff x y (F.2.1 x y f.1)).2 (f.1,<j>F.2.1 x y f.1)@i).1)
(<i> ((ff y x (F.2.1 y x f.2.1)).2 (f.2.1,<j>F.2.1 y x f.2.1)@i).1)
- hole : Id U (iso A x y) (iso B (F.1 x) (F.1 y)) = isoId (iso A x y) (iso B (F.1 x) (F.1 y)) F0 G0 FG GF
+ hole : Path U (iso A x y) (iso B (F.1 x) (F.1 y)) = isoPath (iso A x y) (iso B (F.1 x) (F.1 y)) F0 G0 FG GF
opaque lemFF
F12 (A B : precategory) (isC : isCategory A) (F : functor A B)
(p1 : ffFunctor A B F) (x : cA A) : isContr ((y : cA A) * iso B (F.1 y) (F.1 x))
= transport (<i>isContr ((y : cA A) * (invIsoEq B (F.1 x) (F.1 y)@i))) hole
where
- hole2 (y : cA A) : Id U (iso A x y) (iso B (F.1 x) (F.1 y))
+ hole2 (y : cA A) : Path U (iso A x y) (iso B (F.1 x) (F.1 y))
= lemFF A B F p1 x y
hole : isContr ((y : cA A) * iso B (F.1 x) (F.1 y))
= transport (<i> isContr ((y : cA A) * (hole2 y @ i))) (isC x)
F23 (A B : precategory) (F : functor A B) (p2 : (x : cA A) -> isContr ((y : cA A) * iso B (F.1 y) (F.1 x)))
(x : cA B)
- (a b : (y : cA A) * iso B (F.1 y) x) : Id ((y : cA A) * iso B (F.1 y) x) a b
+ (a b : (y : cA A) * iso B (F.1 y) x) : Path ((y : cA A) * iso B (F.1 y) x) a b
= undefined
-- = transport p hole3
-- where
--- hole2 : Id ((y : cA A) * iso B (F.1 y) (F.1 a.1))
+-- hole2 : Path ((y : cA A) * iso B (F.1 y) (F.1 a.1))
-- (a.1, idIso B (F.1 a.1)) (b.1, compIso B (F.1 b.1) x (F.1 a.1) b.2 (invIso B (F.1 a.1) x a.2))
-- = isContrProp ((y : cA A) * iso B (F.1 y) (F.1 a.1)) (p2 a.1)
-- (a.1, idIso B (F.1 a.1)) (b.1, compIso B (F.1 b.1) x (F.1 a.1) b.2 (invIso B (F.1 a.1) x a.2))
-- opaque hole2
--- hole3 : Id ((y : cA A) * iso B (F.1 y) x)
+-- hole3 : Path ((y : cA A) * iso B (F.1 y) x)
-- (a.1, compIso B (F.1 a.1) (F.1 a.1) x (idIso B (F.1 a.1)) a.2)
-- (b.1, compIso B (F.1 b.1) (F.1 a.1) x (compIso B (F.1 b.1) x (F.1 a.1) b.2 (invIso B (F.1 a.1) x a.2)) a.2)
-- = <i> ((hole2@i).1, compIso B (F.1 (hole2@i).1) (F.1 a.1) x (hole2@i).2 a.2)
-- opaque hole3
--- p : Id U (Id ((y : cA A) * iso B (F.1 y) x)
+-- p : Path U (Path ((y : cA A) * iso B (F.1 y) x)
-- (a.1, compIso B (F.1 a.1) (F.1 a.1) x (idIso B (F.1 a.1)) a.2)
-- (b.1, compIso B (F.1 b.1) (F.1 a.1) x (compIso B (F.1 b.1) x (F.1 a.1) b.2 (invIso B (F.1 a.1) x a.2)) a.2))
--- (Id ((y : cA A) * iso B (F.1 y) x) a b)
--- = (<i>Id ((y : cA A) * iso B (F.1 y) x)
--- (a.1, IdLIso B (F.1 a.1) x a.2 @ i)
--- (b.1, compId (iso B (F.1 b.1) x)
+-- (Path ((y : cA A) * iso B (F.1 y) x) a b)
+-- = (<i>Path ((y : cA A) * iso B (F.1 y) x)
+-- (a.1, PathLIso B (F.1 a.1) x a.2 @ i)
+-- (b.1, compPath (iso B (F.1 b.1) x)
-- (compIso B (F.1 b.1) (F.1 a.1) x (compIso B (F.1 b.1) x (F.1 a.1) b.2 (invIso B (F.1 a.1) x a.2)) a.2)
-- (compIso B (F.1 b.1) x x b.2 (compIso B x (F.1 a.1) x (invIso B (F.1 a.1) x a.2) a.2))
-- b.2
--- (IdCIso B (F.1 b.1) x (F.1 a.1) x b.2 (invIso B (F.1 a.1) x a.2) a.2)
--- (compId (iso B (F.1 b.1) x)
+-- (PathCIso B (F.1 b.1) x (F.1 a.1) x b.2 (invIso B (F.1 a.1) x a.2) a.2)
+-- (compPath (iso B (F.1 b.1) x)
-- (compIso B (F.1 b.1) x x b.2 (compIso B x (F.1 a.1) x (invIso B (F.1 a.1) x a.2) a.2))
-- (compIso B (F.1 b.1) x x b.2 (idIso B x))
-- b.2
--- (<i>compIso B (F.1 b.1) x x b.2 (IdInvLIso B (F.1 a.1) x a.2 @ i))
--- (IdRIso B (F.1 b.1) x b.2))
+-- (<i>compIso B (F.1 b.1) x x b.2 (PathInvLIso B (F.1 a.1) x a.2 @ i))
+-- (PathRIso B (F.1 b.1) x b.2))
-- @ i))
opaque F23
(propFFFunctor A B F)
(\(ff : ffFunctor A B F) -> propPi (cA B) (\(b : cA B) -> (a : cA A) * iso B (F.1 a) b)
(\(b : cA B) -> F23 A B F (F12 A B isC F ff) b))
-catIdIsEquiv (A : precategory) : catIsEquiv A A (idFunctor A)
+catPathIsEquiv (A : precategory) : catIsEquiv A A (idFunctor A)
= (\(a b : cA A) -> idIsEquiv (cH A a b)
,\(b:cA A) -> (b, idIso A b))
-catIdEquiv (A : precategory) : catEquiv A A = (idFunctor A, catIdIsEquiv A)
+catPathEquiv (A : precategory) : catEquiv A A = (idFunctor A, catPathIsEquiv A)
catIsIso (A B : precategory) (F : functor A B) : U
= (_ : ffFunctor A B F) * isEquiv (cA A) (cA B) F.1
catIso21 (A B : precategory) (e1:equiv (cA A) (cA B)) : U
= (x y:cA A) -> equiv (cH A x y) (cH B (e1.1 x) (e1.1 y))
catIso22 (A B : precategory) (e1:equiv (cA A) (cA B)) (e2:catIso21 A B e1) : U
- = (x:cA A) -> Id (cH B (e1.1 x) (e1.1 x))
- ((e2 x x).1 (cId A x))
- (cId B (e1.1 x))
+ = (x:cA A) -> Path (cH B (e1.1 x) (e1.1 x))
+ ((e2 x x).1 (cPath A x))
+ (cPath B (e1.1 x))
catIso23 (A B : precategory) (e1:equiv (cA A) (cA B)) (e2:catIso21 A B e1) : U
= (x y z:cA A)(f:cH A x y)(g:cH A y z)->
- Id (cH B (e1.1 x) (e1.1 z))
+ Path (cH B (e1.1 x) (e1.1 z))
((e2 x z).1 (cC A x y z f g))
(cC B (e1.1 x) (e1.1 y) (e1.1 z) ((e2 x y).1 f) ((e2 y z).1 g))
catIso2 (A B : precategory) : U = (e1 : equiv (cA A) (cA B))
* (_ : catIso22 A B e1 e2)
* ( catIso23 A B e1 e2)
-eCatIso (A B : precategory) : Id U (catIso A B) (catIso2 A B)
- = transEquivToId (catIso A B) (catIso2 A B)
+eCatIso (A B : precategory) : Path U (catIso A B) (catIso2 A B)
+ = transEquivToPath (catIso A B) (catIso2 A B)
(F, gradLemma (catIso A B) (catIso2 A B) F G (\(e:catIso2 A B)-><_>e) (\(e:catIso A B)-><_>e))
where
F(e:catIso A B):catIso2 A B=((e.1.1, e.2.2),\(x y:cA A)->(e.1.2.1 x y, e.2.1 x y),e.1.2.2.1,e.1.2.2.2)
G(e:catIso2 A B):catIso A B=((e.1.1,\(x y:cA A)->(e.2.1 x y).1,e.2.2.1,e.2.2.2),\(x y:cA A)->(e.2.1 x y).2,e.1.2)
-catIso21' (A B : precategory) (e1:Id U (cA A) (cA B)) : U
+catIso21' (A B : precategory) (e1:Path U (cA A) (cA B)) : U
= (x y:cA A) -> equiv (cH A x y) (cH B (transport e1 x) (transport e1 y))
-catIso22' (A B : precategory) (e1:Id U (cA A) (cA B)) (e2:catIso21' A B e1) : U
- = (x:cA A) -> Id (cH B (transport e1 x) (transport e1 x))
- ((e2 x x).1 (cId A x))
- (cId B (transport e1 x))
-catIso23' (A B : precategory) (e1:Id U (cA A) (cA B)) (e2:catIso21' A B e1) : U
+catIso22' (A B : precategory) (e1:Path U (cA A) (cA B)) (e2:catIso21' A B e1) : U
+ = (x:cA A) -> Path (cH B (transport e1 x) (transport e1 x))
+ ((e2 x x).1 (cPath A x))
+ (cPath B (transport e1 x))
+catIso23' (A B : precategory) (e1:Path U (cA A) (cA B)) (e2:catIso21' A B e1) : U
= (x y z:cA A)(f:cH A x y)(g:cH A y z)->
- Id (cH B (transport e1 x) (transport e1 z))
+ Path (cH B (transport e1 x) (transport e1 z))
((e2 x z).1 (cC A x y z f g))
(cC B (transport e1 x) (transport e1 y) (transport e1 z) ((e2 x y).1 f) ((e2 y z).1 g))
)
where F(f:(a:A)->B0 a)(a:A):B1 a= (e a).1 (f a)
-eCatIso2 (A B:precategory):Id U (catIso32 A B) (catIso2 A B)
- = sigEquivLem' (Id U (cA A) (cA B)) (equiv (cA A) (cA B))
- (\(e1:Id U (cA A) (cA B)) -> (e2:catIso321 A B e1)*(_:catIso322 A B e1 e2)*(catIso323 A B e1 e2))
+eCatIso2 (A B:precategory):Path U (catIso32 A B) (catIso2 A B)
+ = sigEquivLem' (Path U (cA A) (cA B)) (equiv (cA A) (cA B))
+ (\(e1:Path U (cA A) (cA B)) -> (e2:catIso321 A B e1)*(_:catIso322 A B e1 e2)*(catIso323 A B e1 e2))
(\(e1:equiv (cA A) (cA B)) -> (e2:catIso21 A B e1)*(_:catIso22 A B e1 e2)*(catIso23 A B e1 e2))
(corrUniv' (cA A) (cA B))
- (\(e1:Id U (cA A) (cA B)) -> let e1' : equiv (cA A) (cA B) = transEquiv' (cA B) (cA A) e1
- p0 (x:cA A):Id (cA B) (transport e1 x) (e1'.1 x)=lemTransEquiv (cA A) (cA B) e1 x
+ (\(e1:Path U (cA A) (cA B)) -> let e1' : equiv (cA A) (cA B) = transEquiv' (cA B) (cA A) e1
+ p0 (x:cA A):Path (cA B) (transport e1 x) (e1'.1 x)=lemTransEquiv (cA A) (cA B) e1 x
in
transport
- (<i> Id U ((e2:catIso321 A B e1)*(_:catIso322 A B e1 e2)*(catIso323 A B e1 e2))
+ (<i> Path U ((e2:catIso321 A B e1)*(_:catIso322 A B e1 e2)*(catIso323 A B e1 e2))
((e2 : (x y:cA A) -> equiv (cH A x y) (cH B (p0 x@i) (p0 y@i)))
- * (_ : (x:cA A) -> Id (cH B (p0 x@i) (p0 x@i)) ((e2 x x).1 (cId A x)) (cId B (p0 x@i)))
+ * (_ : (x:cA A) -> Path (cH B (p0 x@i) (p0 x@i)) ((e2 x x).1 (cPath A x)) (cPath B (p0 x@i)))
* ((x y z:cA A)(f:cH A x y)(g:cH A y z)->
- Id (cH B (p0 x@i) (p0 z@i))
+ Path (cH B (p0 x@i) (p0 z@i))
((e2 x z).1 (cC A x y z f g))
(cC B (p0 x@i) (p0 y@i) (p0 z@i) ((e2 x y).1 f) ((e2 y z).1 g))))
)
(\(e2:catIso321 A B e1) -> (_:catIso322 A B e1 e2)*(catIso323 A B e1 e2))
(\(e2:catIso21' A B e1) -> (_:catIso22' A B e1 e2)*(catIso23' A B e1 e2))
(equivPi (cA A)
- (\(x:cA A) -> (y:cA A) -> Id U (cH A x y) (cH B (transport e1 x) (transport e1 y)))
+ (\(x:cA A) -> (y:cA A) -> Path U (cH A x y) (cH B (transport e1 x) (transport e1 y)))
(\(x:cA A) -> (y:cA A) -> equiv (cH A x y) (cH B (transport e1 x) (transport e1 y)))
(\(x:cA A) -> equivPi (cA A)
- (\(y:cA A) -> Id U (cH A x y) (cH B (transport e1 x) (transport e1 y)))
+ (\(y:cA A) -> Path U (cH A x y) (cH B (transport e1 x) (transport e1 y)))
(\(y:cA A) -> equiv (cH A x y) (cH B (transport e1 x) (transport e1 y)))
(\(y:cA A) -> corrUniv' (cH A x y) (cH B (transport e1 x) (transport e1 y)))))
(\(e2:catIso321 A B e1) -> let e2' (x y:cA A) : equiv (cH A x y) (cH B (transport e1 x) (transport e1 y))
= transEquiv' (cH B (transport e1 x) (transport e1 y)) (cH A x y) (e2 x y)
- p2 (x y:cA A)(f:cH A x y):Id (cH B (transport e1 x) (transport e1 y)) (transport (e2 x y) f) ((e2' x y).1 f)
+ p2 (x y:cA A)(f:cH A x y):Path (cH B (transport e1 x) (transport e1 y)) (transport (e2 x y) f) ((e2' x y).1 f)
= lemTransEquiv (cH A x y) (cH B (transport e1 x) (transport e1 y)) (e2 x y) f
in
- <i> (_:(x:cA A) -> Id (cH B (transport e1 x) (transport e1 x)) (p2 x x (cId A x)@i) (cId B (transport e1 x)))
+ <i> (_:(x:cA A) -> Path (cH B (transport e1 x) (transport e1 x)) (p2 x x (cPath A x)@i) (cPath B (transport e1 x)))
* ((x y z:cA A)(f:cH A x y)(g:cH A y z)->
- Id (cH B (transport e1 x) (transport e1 z))
+ Path (cH B (transport e1 x) (transport e1 z))
(p2 x z (cC A x y z f g)@i)
(cC B (transport e1 x) (transport e1 y) (transport e1 z) (p2 x y f@i) (p2 y z g@i)))
)))
catIsoIsEquiv (A B : precategory) (F : functor A B) (e : catIsIso A B F) : catIsEquiv A B F
= (e.1,\(b:cA B)->((e.2 b).1.1, eqToIso B (F.1 (e.2 b).1.1) b (<i>(e.2 b).1.2@-i)))
-invEquiv (A:U) (a b:A) : Id U (Id A a b) (Id A b a)
- = equivId (Id A a b) (Id A b a) (inv A a b) (gradLemma (Id A a b) (Id A b a) (inv A a b) (inv A b a) (\(x:Id A b a) -> <_> x) (\(x:Id A a b) -> <_> x))
+invEquiv (A:U) (a b:A) : Path U (Path A a b) (Path A b a)
+ = equivPath (Path A a b) (Path A b a) (inv A a b) (gradLemma (Path A a b) (Path A b a) (inv A a b) (inv A b a) (\(x:Path A b a) -> <_> x) (\(x:Path A a b) -> <_> x))
catEquivIsIso (A B : precategory) (isCA : isCategory A) (isCB : isCategory B) (F : functor A B) (e : catIsEquiv A B F)
: catIsIso A B F
= equivProp (catIsEquiv A B F) (catIsIso A B F) (catPropIsEquiv A B isCA F) (catPropIsIso A B F)
(catEquivIsIso A B isCA isCB F) (catIsoIsEquiv A B F)
-catEquivEqIso (A B : precategory) (isCA : isCategory A) (isCB : isCategory B) : Id U (catEquiv A B) (catIso A B)
- = <i> (F : functor A B) * (transEquivToId (catIsEquiv A B F) (catIsIso A B F) (catIsEquivEqIso A B isCA isCB F) @ i)
+catEquivEqIso (A B : precategory) (isCA : isCategory A) (isCB : isCategory B) : Path U (catEquiv A B) (catIso A B)
+ = <i> (F : functor A B) * (transEquivToPath (catIsEquiv A B F) (catIsIso A B F) (catIsEquivEqIso A B isCA isCB F) @ i)
-catIsoEqId (A B : precategory) : Id U (catIso A B) (Id precategory A B)
- = compId U (catIso A B) (catIso2 A B) (Id precategory A B)
+catIsoEqPath (A B : precategory) : Path U (catIso A B) (Path precategory A B)
+ = compPath U (catIso A B) (catIso2 A B) (Path precategory A B)
(eCatIso A B)
- (compId U (catIso2 A B) (catIso32 A B) (Id precategory A B)
+ (compPath U (catIso2 A B) (catIso32 A B) (Path precategory A B)
(<i>eCatIso2 A B@-i)
- (compId U (catIso32 A B) (catIso31 A B) (Id precategory A B)
+ (compPath U (catIso32 A B) (catIso31 A B) (Path precategory A B)
(<i>eCatIso32 A B@-i)
- (compId U (catIso31 A B) (catIso30 A B) (Id precategory A B)
+ (compPath U (catIso31 A B) (catIso30 A B) (Path precategory A B)
(<i>eCatIso31 A B@-i)
- (compId U (catIso30 A B) (catIso3 A B) (Id precategory A B)
+ (compPath U (catIso30 A B) (catIso3 A B) (Path precategory A B)
(<i>eCatIso30 A B@-i)
(<i>eCatIso3 A B@-i)))))
-catEquivEqId (A B : precategory) (isCA : isCategory A) (isCB : isCategory B) : Id U (catEquiv A B) (Id precategory A B)
- = compId U (catEquiv A B) (catIso A B) (Id precategory A B) (catEquivEqIso A B isCA isCB) (catIsoEqId A B)
+catEquivEqPath (A B : precategory) (isCA : isCategory A) (isCB : isCategory B) : Path U (catEquiv A B) (Path precategory A B)
+ = compPath U (catEquiv A B) (catIso A B) (Path precategory A B) (catEquivEqIso A B isCA isCB) (catIsoEqPath A B)
-catEquivEqId' (A : category) : isContr ((B : category) * catEquiv A.1 B.1)
- = transport (<i> isContr ((B : category) * catEquivEqId A.1 B.1 A.2 B.2@-i))
+catEquivEqPath' (A : category) : isContr ((B : category) * catEquiv A.1 B.1)
+ = transport (<i> isContr ((B : category) * catEquivEqPath A.1 B.1 A.2 B.2@-i))
((A, <_> A.1)
- ,\(BB:((B : category) * Id precategory A.1 B.1))->
+ ,\(BB:((B : category) * Path precategory A.1 B.1))->
<i>((BB.2@i
- ,lemIdPProp (isCategory A.1)
+ ,lemPathPProp (isCategory A.1)
(isCategory BB.1.1)
(propIsCategory A.1)
(<i> isCategory (BB.2@i))
= (X : cA C) * (_ : homTo C X) * homTo C X
hasCospanCone (C : precategory) (D : cospan C) (W : cA C) : U
= (f : cH C W D.2.1.1) * (g : cH C W D.2.2.1)
- * Id (cH C W D.1)
+ * Path (cH C W D.1)
(cC C W D.2.1.1 D.1 f D.2.1.2)
(cC C W D.2.2.1 D.1 g D.2.2.2)
cospanCone (C : precategory) (D : cospan C) : U = (W : cA C) * hasCospanCone C D W
isCospanConeHom (C : precategory) (D : cospan C) (E1 E2 : cospanCone C D) (h : cH C E1.1 E2.1) : U
- = (_ : Id (cH C E1.1 D.2.1.1) (cC C E1.1 E2.1 D.2.1.1 h E2.2.1) E1.2.1)
- * (Id (cH C E1.1 D.2.2.1) (cC C E1.1 E2.1 D.2.2.1 h E2.2.2.1) E1.2.2.1)
+ = (_ : Path (cH C E1.1 D.2.1.1) (cC C E1.1 E2.1 D.2.1.1 h E2.2.1) E1.2.1)
+ * (Path (cH C E1.1 D.2.2.1) (cC C E1.1 E2.1 D.2.2.1 h E2.2.2.1) E1.2.2.1)
isCospanConeHomProp (C : precategory) (D : cospan C) (E1 E2 : cospanCone C D) (h : cH C E1.1 E2.1)
: prop (isCospanConeHom C D E1 E2 h)
- = propAnd (Id (cH C E1.1 D.2.1.1) (cC C E1.1 E2.1 D.2.1.1 h E2.2.1) E1.2.1)
- (Id (cH C E1.1 D.2.2.1) (cC C E1.1 E2.1 D.2.2.1 h E2.2.2.1) E1.2.2.1)
+ = propAnd (Path (cH C E1.1 D.2.1.1) (cC C E1.1 E2.1 D.2.1.1 h E2.2.1) E1.2.1)
+ (Path (cH C E1.1 D.2.2.1) (cC C E1.1 E2.1 D.2.2.1 h E2.2.2.1) E1.2.2.1)
(cHSet C E1.1 D.2.1.1 (cC C E1.1 E2.1 D.2.1.1 h E2.2.1) E1.2.1)
(cHSet C E1.1 D.2.2.1 (cC C E1.1 E2.1 D.2.2.1 h E2.2.2.1) E1.2.2.1)
cospanConeHom (C : precategory) (D : cospan C) (E1 E2 : cospanCone C D) : U
= (h : cH C E1.1 E2.1) * isCospanConeHom C D E1 E2 h
-cospanConeId (C : precategory) (D : cospan C) (E : cospanCone C D) : cospanConeHom C D E E
- = (cId C E.1, cIdL C E.1 D.2.1.1 E.2.1, cIdL C E.1 D.2.2.1 E.2.2.1)
+cospanConePath (C : precategory) (D : cospan C) (E : cospanCone C D) : cospanConeHom C D E E
+ = (cPath C E.1, cPathL C E.1 D.2.1.1 E.2.1, cPathL C E.1 D.2.2.1 E.2.2.1)
cospanConeComp (C : precategory) (D : cospan C) (X Y Z : cospanCone C D)
(F : cospanConeHom C D X Y) (G : cospanConeHom C D Y Z) : cospanConeHom C D X Z
= (cC C X.1 Y.1 Z.1 F.1 G.1
- ,compId (cH C X.1 D.2.1.1)
+ ,compPath (cH C X.1 D.2.1.1)
(cC C X.1 Z.1 D.2.1.1 (cC C X.1 Y.1 Z.1 F.1 G.1) Z.2.1)
(cC C X.1 Y.1 D.2.1.1 F.1 (cC C Y.1 Z.1 D.2.1.1 G.1 Z.2.1))
X.2.1
- (cIdC C X.1 Y.1 Z.1 D.2.1.1 F.1 G.1 Z.2.1)
- (compId (cH C X.1 D.2.1.1)
+ (cPathC C X.1 Y.1 Z.1 D.2.1.1 F.1 G.1 Z.2.1)
+ (compPath (cH C X.1 D.2.1.1)
(cC C X.1 Y.1 D.2.1.1 F.1 (cC C Y.1 Z.1 D.2.1.1 G.1 Z.2.1))
(cC C X.1 Y.1 D.2.1.1 F.1 Y.2.1)
X.2.1
(<i> cC C X.1 Y.1 D.2.1.1 F.1 (G.2.1 @ i))
F.2.1)
- ,compId (cH C X.1 D.2.2.1)
+ ,compPath (cH C X.1 D.2.2.1)
(cC C X.1 Z.1 D.2.2.1 (cC C X.1 Y.1 Z.1 F.1 G.1) Z.2.2.1)
(cC C X.1 Y.1 D.2.2.1 F.1 (cC C Y.1 Z.1 D.2.2.1 G.1 Z.2.2.1))
X.2.2.1
- (cIdC C X.1 Y.1 Z.1 D.2.2.1 F.1 G.1 Z.2.2.1)
- (compId (cH C X.1 D.2.2.1)
+ (cPathC C X.1 Y.1 Z.1 D.2.2.1 F.1 G.1 Z.2.2.1)
+ (compPath (cH C X.1 D.2.2.1)
(cC C X.1 Y.1 D.2.2.1 F.1 (cC C Y.1 Z.1 D.2.2.1 G.1 Z.2.2.1))
(cC C X.1 Y.1 D.2.2.1 F.1 Y.2.2.1)
X.2.2.1
= (hasCospanCone C D
,\(x y : cA C) (a : hasCospanCone C D x) (b : hasCospanCone C D y) -> isCospanConeHom C D (x, a) (y, b)
,\(x y : cA C) (a : hasCospanCone C D x) (b : hasCospanCone C D y) -> isCospanConeHomProp C D (x, a) (y, b)
- ,\(x : cA C) (a : hasCospanCone C D x) -> (cospanConeId C D (x, a)).2
+ ,\(x : cA C) (a : hasCospanCone C D x) -> (cospanConePath C D (x, a)).2
,\(x y z : cA C) (a : hasCospanCone C D x) (b : hasCospanCone C D y) (c : hasCospanCone C D z)
(f : cH C x y) (g : cH C y z)
(Hf : isCospanConeHom C D (x, a) (y, b) f)
where
hole : isStandardStructure C (cospanConeStructure C D)
= \(x : cA C) (a b : hasCospanCone C D x)
- (c : isCospanConeHom C D (x, a) (x, b) (cId C x))
- (d : isCospanConeHom C D (x, b) (x, a) (cId C x)) ->
- <i> (compId (cH C x D.2.1.1) a.1 (cC C x x D.2.1.1 (cId C x) a.1) b.1 (<i>cIdL C x D.2.1.1 a.1 @-i) d.1 @ i
- ,compId (cH C x D.2.2.1) a.2.1 (cC C x x D.2.2.1 (cId C x) a.2.1) b.2.1 (<i>cIdL C x D.2.2.1 a.2.1 @-i) d.2 @ i
- ,lemIdPProp (Id (cH C x D.1) (cC C x D.2.1.1 D.1 a.1 D.2.1.2) (cC C x D.2.2.1 D.1 a.2.1 D.2.2.2))
- (Id (cH C x D.1) (cC C x D.2.1.1 D.1 b.1 D.2.1.2) (cC C x D.2.2.1 D.1 b.2.1 D.2.2.2))
+ (c : isCospanConeHom C D (x, a) (x, b) (cPath C x))
+ (d : isCospanConeHom C D (x, b) (x, a) (cPath C x)) ->
+ <i> (compPath (cH C x D.2.1.1) a.1 (cC C x x D.2.1.1 (cPath C x) a.1) b.1 (<i>cPathL C x D.2.1.1 a.1 @-i) d.1 @ i
+ ,compPath (cH C x D.2.2.1) a.2.1 (cC C x x D.2.2.1 (cPath C x) a.2.1) b.2.1 (<i>cPathL C x D.2.2.1 a.2.1 @-i) d.2 @ i
+ ,lemPathPProp (Path (cH C x D.1) (cC C x D.2.1.1 D.1 a.1 D.2.1.2) (cC C x D.2.2.1 D.1 a.2.1 D.2.2.2))
+ (Path (cH C x D.1) (cC C x D.2.1.1 D.1 b.1 D.2.1.2) (cC C x D.2.2.1 D.1 b.2.1 D.2.2.2))
(cHSet C x D.1 (cC C x D.2.1.1 D.1 a.1 D.2.1.2) (cC C x D.2.2.1 D.1 a.2.1 D.2.2.2))
- (<i>Id (cH C x D.1)
- (cC C x D.2.1.1 D.1 (compId (cH C x D.2.1.1) a.1 (cC C x x D.2.1.1 (cId C x) a.1) b.1
- (<i>cIdL C x D.2.1.1 a.1 @-i) d.1 @ i) D.2.1.2)
- (cC C x D.2.2.1 D.1 (compId (cH C x D.2.2.1) a.2.1 (cC C x x D.2.2.1 (cId C x) a.2.1) b.2.1
- (<i>cIdL C x D.2.2.1 a.2.1 @-i) d.2 @ i) D.2.2.2))
+ (<i>Path (cH C x D.1)
+ (cC C x D.2.1.1 D.1 (compPath (cH C x D.2.1.1) a.1 (cC C x x D.2.1.1 (cPath C x) a.1) b.1
+ (<i>cPathL C x D.2.1.1 a.1 @-i) d.1 @ i) D.2.1.2)
+ (cC C x D.2.2.1 D.1 (compPath (cH C x D.2.2.1) a.2.1 (cC C x x D.2.2.1 (cPath C x) a.2.1) b.2.1
+ (<i>cPathL C x D.2.2.1 a.2.1 @-i) d.2 @ i) D.2.2.2))
a.2.2 b.2.2 @ i)
isFinal (C : precategory) (A : cA C) : U = (B : cA C) -> isContr (cH C B A)
= \(x y : hasFinal C) ->
let p : iso C x.1 y.1
= ((y.2 x.1).1, (x.2 y.1).1
- , isContrProp (cH C x.1 x.1) (x.2 x.1) (cC C x.1 y.1 x.1 (y.2 x.1).1 (x.2 y.1).1) (cId C x.1)
- , isContrProp (cH C y.1 y.1) (y.2 y.1) (cC C y.1 x.1 y.1 (x.2 y.1).1 (y.2 x.1).1) (cId C y.1))
+ , isContrProp (cH C x.1 x.1) (x.2 x.1) (cC C x.1 y.1 x.1 (y.2 x.1).1 (x.2 y.1).1) (cPath C x.1)
+ , isContrProp (cH C y.1 y.1) (y.2 y.1) (cC C y.1 x.1 y.1 (x.2 y.1).1 (y.2 x.1).1) (cPath C y.1))
in <i> ( (lemIsCategory C isC x.1 y.1 p @ i).1
- , lemIdPProp (isFinal C x.1)
+ , lemPathPProp (isFinal C x.1)
(isFinal C y.1)
(isFinalProp C x.1)
(<i> isFinal C (lemIsCategory C isC x.1 y.1 p @ i).1)
(A B C D : cA CA)
(f : cH CA A B) (g : cH CA C D)
(h : cH CA A C) (i : cH CA B D)
- : U = Id (cH CA A D) (cC CA A C D h g) (cC CA A B D f i)
+ : U = Path (cH CA A D) (cC CA A C D h g) (cC CA A B D f i)
pullbackPasting
(CA : precategory)
(pb3 : isPullback CA (F, (D, cC CA D E F h i), (C, l)) (A, j, cC CA A B C f g, cc3))
: isPullback CA (E, (D, h), (B, k)) (A, j, f, cc1)
= \(c : cospanCone CA (E, (D, h), (B, k))) ->
- let hole : Id (cH CA c.1 F)
+ let hole : Path (cH CA c.1 F)
(cC CA c.1 D F c.2.1 (cC CA D E F h i))
(cC CA c.1 C F (cC CA c.1 B C c.2.2.1 g) l)
- = compId (cH CA c.1 F)
+ = compPath (cH CA c.1 F)
(cC CA c.1 D F c.2.1 (cC CA D E F h i))
(cC CA c.1 E F (cC CA c.1 D E c.2.1 h) i)
(cC CA c.1 C F (cC CA c.1 B C c.2.2.1 g) l)
- (<n>cIdC CA c.1 D E F c.2.1 h i@-n)
- (compId (cH CA c.1 F)
+ (<n>cPathC CA c.1 D E F c.2.1 h i@-n)
+ (compPath (cH CA c.1 F)
(cC CA c.1 E F (cC CA c.1 D E c.2.1 h) i)
(cC CA c.1 E F (cC CA c.1 B E c.2.2.1 k) i)
(cC CA c.1 C F (cC CA c.1 B C c.2.2.1 g) l)
(<n>cC CA c.1 E F (c.2.2.2@n) i)
- (compId (cH CA c.1 F)
+ (compPath (cH CA c.1 F)
(cC CA c.1 E F (cC CA c.1 B E c.2.2.1 k) i)
(cC CA c.1 B F c.2.2.1 (cC CA B E F k i))
(cC CA c.1 C F (cC CA c.1 B C c.2.2.1 g) l)
- (cIdC CA c.1 B E F c.2.2.1 k i)
- (compId (cH CA c.1 F)
+ (cPathC CA c.1 B E F c.2.2.1 k i)
+ (compPath (cH CA c.1 F)
(cC CA c.1 B F c.2.2.1 (cC CA B E F k i))
(cC CA c.1 B F c.2.2.1 (cC CA B C F g l))
(cC CA c.1 C F (cC CA c.1 B C c.2.2.1 g) l)
(<n>cC CA c.1 B F c.2.2.1 (cc2@n))
- (<n>cIdC CA c.1 B C F c.2.2.1 g l@-n))))
+ (<n>cPathC CA c.1 B C F c.2.2.1 g l@-n))))
c' : cospanCone CA (F, (D, cC CA D E F h i), (C, l))
= (c.1
,c.2.1
,cC CA c.1 B C c.2.2.1 g
,hole)
- hole2 : Id (cH CA c.1 F)
+ hole2 : Path (cH CA c.1 F)
(cC CA c.1 E F (cC CA c.1 D E c.2.1 h) i)
(cC CA c.1 C F (cC CA c.1 B C c.2.2.1 g) l)
- = compId (cH CA c.1 F)
+ = compPath (cH CA c.1 F)
(cC CA c.1 E F (cC CA c.1 D E c.2.1 h) i)
(cC CA c.1 D F c.2.1 (cC CA D E F h i))
(cC CA c.1 C F (cC CA c.1 B C c.2.2.1 g) l)
- (cIdC CA c.1 D E F c.2.1 h i)
+ (cPathC CA c.1 D E F c.2.1 h i)
hole
cc : cospanCone CA (F, (E, i), (C, l))
= (c.1, cC CA c.1 D E c.2.1 h, c'.2.2.1, hole2)
- p0 (h' : cH CA c.1 A) (p : Id (cH CA c.1 D) (cC CA c.1 A D h' j) c.2.1)
- : Id U (Id (cH CA c.1 B) (cC CA c.1 A B h' f) c.2.2.1)
- (Id (cH CA c.1 C) (cC CA c.1 A C h' (cC CA A B C f g)) c'.2.2.1)
+ p0 (h' : cH CA c.1 A) (p : Path (cH CA c.1 D) (cC CA c.1 A D h' j) c.2.1)
+ : Path U (Path (cH CA c.1 B) (cC CA c.1 A B h' f) c.2.2.1)
+ (Path (cH CA c.1 C) (cC CA c.1 A C h' (cC CA A B C f g)) c'.2.2.1)
= transport
- (<i> Id U (Id (cH CA c.1 B) (cC CA c.1 A B h' f) c.2.2.1)
- (Id (cH CA c.1 C) (cIdC CA c.1 A B C h' f g @ i) c'.2.2.1))
- (idProp (Id (cH CA c.1 B) (cC CA c.1 A B h' f) c.2.2.1)
- (Id (cH CA c.1 C) (cIdC CA c.1 A B C h' f g @ 0) c'.2.2.1)
+ (<i> Path U (Path (cH CA c.1 B) (cC CA c.1 A B h' f) c.2.2.1)
+ (Path (cH CA c.1 C) (cPathC CA c.1 A B C h' f g @ i) c'.2.2.1))
+ (idProp (Path (cH CA c.1 B) (cC CA c.1 A B h' f) c.2.2.1)
+ (Path (cH CA c.1 C) (cPathC CA c.1 A B C h' f g @ 0) c'.2.2.1)
(cHSet CA c.1 B (cC CA c.1 A B h' f) c.2.2.1)
- (cHSet CA c.1 C (cIdC CA c.1 A B C h' f g @ 0) c'.2.2.1)
- (\(p:Id (cH CA c.1 B) (cC CA c.1 A B h' f) c.2.2.1) ->
+ (cHSet CA c.1 C (cPathC CA c.1 A B C h' f g @ 0) c'.2.2.1)
+ (\(p:Path (cH CA c.1 B) (cC CA c.1 A B h' f) c.2.2.1) ->
<i> cC CA c.1 B C (p@i) g)
- (\(p1:Id (cH CA c.1 C) (cIdC CA c.1 A B C h' f g @ 0) c'.2.2.1) ->
+ (\(p1:Path (cH CA c.1 C) (cPathC CA c.1 A B C h' f g @ 0) c'.2.2.1) ->
let h1 : cospanConeHom CA (F, (E, i), (C, l)) cc (B, k, g, cc2)
= (cC CA c.1 A B h' f
- ,compId (cH CA c.1 E)
+ ,compPath (cH CA c.1 E)
(cC CA c.1 B E (cC CA c.1 A B h' f) k)
(cC CA c.1 A E h' (cC CA A B E f k))
(cC CA c.1 D E c.2.1 h)
- (cIdC CA c.1 A B E h' f k)
- (compId (cH CA c.1 E)
+ (cPathC CA c.1 A B E h' f k)
+ (compPath (cH CA c.1 E)
(cC CA c.1 A E h' (cC CA A B E f k))
(cC CA c.1 A E h' (cC CA A D E j h))
(cC CA c.1 D E c.2.1 h)
(<i>cC CA c.1 A E h' (cc1@-i))
- (compId (cH CA c.1 E)
+ (compPath (cH CA c.1 E)
(cC CA c.1 A E h' (cC CA A D E j h))
(cC CA c.1 D E (cC CA c.1 A D h' j) h)
(cC CA c.1 D E c.2.1 h)
- (<i> cIdC CA c.1 A D E h' j h @ -i)
+ (<i> cPathC CA c.1 A D E h' j h @ -i)
(<i> cC CA c.1 D E (p@i) h)))
,p1)
h2 : cospanConeHom CA (F, (E, i), (C, l)) cc (B, k, g, cc2)
in <n> (isContrProp (cospanConeHom CA (F, (E, i), (C, l)) cc (B, k, g, cc2)) (pb2 cc) h1 h2 @ n).1
))
- p : Id U (cospanConeHom CA (F, (D, cC CA D E F h i), (C, l)) c' (A, j, cC CA A B C f g, cc3))
+ p : Path U (cospanConeHom CA (F, (D, cC CA D E F h i), (C, l)) c' (A, j, cC CA A B C f g, cc3))
(cospanConeHom CA (E, (D, h), (B, k)) c (A, j, f, cc1))
= <i> (h : cH CA c.1 A)
- * (p : Id (cH CA c.1 D) (cC CA c.1 A D h j) c.2.1)
+ * (p : Path (cH CA c.1 D) (cC CA c.1 A D h j) c.2.1)
* (p0 h p @ -i)
in transport (<i> isContr (p@i)) (pb3 c')
data S1 = base
| loop <i> [(i=0) -> base, (i=1) -> base]
-loopS1 : U = Id S1 base base
+loopS1 : U = Path S1 base base
loop1 : loopS1 = <i> loop{S1} @ i
helix : S1 -> U = split
base -> Z
- loop @ i -> sucIdZ @ i
+ loop @ i -> sucPathZ @ i
winding (p : loopS1) : Z = trans Z Z rem zeroZ
where
- rem : Id U Z Z = <i> helix (p @ i)
+ rem : Path U Z Z = <i> helix (p @ i)
-compS1 : loopS1 -> loopS1 -> loopS1 = compId S1 base base base
+compS1 : loopS1 -> loopS1 -> loopS1 = compPath S1 base base base
-- All of these should be equal to "posZ (suc zero)":
loop2 : loopS1 = compS1 loop1 loop1
-loop2' : loopS1 = compId' S1 base base base loop1 loop1
-loop2'' : loopS1 = compId'' S1 base base loop1 base loop1
+loop2' : loopS1 = compPath' S1 base base base loop1 loop1
+loop2'' : loopS1 = compPath'' S1 base base loop1 base loop1
-- More examples:
loopZ1 : Z = winding loop1
loopZN1 : Z = winding invLoop
loopZ0 : Z = winding (compS1 loop1 invLoop)
-mLoop : (x : S1) -> Id S1 x x = split
+mLoop : (x : S1) -> Path S1 x x = split
base -> loop1
loop @ i -> constSquare S1 base loop1 @ i
-- A nice example of a homotopy on the circle. The path going halfway
-- around the circle and then back is contractible:
-hmtpy : Id loopS1 (<i> base) (<i> loop{S1} @ (i /\ -i)) =
+hmtpy : Path loopS1 (<i> base) (<i> loop{S1} @ (i /\ -i)) =
<j i> loop{S1} @ j /\ i /\ -i
-circleelim (X : U) (x : X) (p : Id X x x) : S1 -> X = split
+circleelim (X : U) (x : X) (p : Path X x x) : S1 -> X = split
base -> x
loop @ i -> p @ i
-apcircleelim (A B : U) (x : A) (p : Id A x x) (f : A -> B) :
- (z : S1) -> Id B (f (circleelim A x p z))
+apcircleelim (A B : U) (x : A) (p : Path A x x) (f : A -> B) :
+ (z : S1) -> Path B (f (circleelim A x p z))
(circleelim B (f x) (<i> f (p @ i)) z) = split
base -> <_> f x
loop @ i -> <_> f (p @ i)
-- a special case, Lemmas 6.2.5-6.2.9 in the book
-aLoop (A:U) : U = (a:A) * Id A a a
+aLoop (A:U) : U = (a:A) * Path A a a
phi (A:U) (al : aLoop A) : S1 -> A = split
base -> al.1
psi (A:U) (f:S1 -> A) : aLoop A = (f base,<i>f (loop1@i))
-rem (A:U) (f : S1 -> A) : (u : S1) -> Id A (phi A (psi A f) u) (f u) = split
+rem (A:U) (f : S1 -> A) : (u : S1) -> Path A (phi A (psi A f) u) (f u) = split
base -> refl A (f base)
loop @ i -> <j>f (loop1@i)
-lem (A:U) (f : S1 -> A) : Id (S1 -> A) (phi A (psi A f)) f =
+lem (A:U) (f : S1 -> A) : Path (S1 -> A) (phi A (psi A f)) f =
<i> \ (x:S1) -> (rem A f x) @ i
-thm (A:U) : Id U (aLoop A) (S1 -> A) = isoId T0 T1 f g t s
+thm (A:U) : Path U (aLoop A) (S1 -> A) = isoPath T0 T1 f g t s
where T0 : U = aLoop A
T1 : U = S1 -> A
f : T0 -> T1 = phi A
g : T1 -> T0 = psi A
- s (x:T0) : Id T0 (g (f x)) x = refl T0 x
- t : (y:T1) -> Id T1 (f (g y)) y = lem A
+ s (x:T0) : Path T0 (g (f x)) x = refl T0 x
+ t : (y:T1) -> Path T1 (f (g y)) y = lem A
setPi A (\(x : A) -> B) (\(x : A) -> sB)
eqEquivFst (A B : U) : (t u : equiv A B) ->
- Id U (Id (equiv A B) t u) (Id (A -> B) t.1 u.1)
+ Path U (Path (equiv A B) t u) (Path (A -> B) t.1 u.1)
= lemSigProp (A -> B) (isEquiv A B) (propIsEquiv A B)
-- groupoidFun (A B : U) (gB:groupoid B) : groupoid (A -> B) =
-- groupoidPi A (\(x : A) -> B) (\(x : A) -> gB)
--- lem5 (A B : U) (gB:groupoid B) (t u:equiv A B) : set (Id (equiv A B) t u)
--- = substInv U set (Id (equiv A B) t u) (Id (A -> B) t.1 u.1)
+-- lem5 (A B : U) (gB:groupoid B) (t u:equiv A B) : set (Path (equiv A B) t u)
+-- = substInv U set (Path (equiv A B) t u) (Path (A -> B) t.1 u.1)
-- (eqEquivFst A B t u) (groupoidFun A B gB t.1 u.1)
-setId (A B : U) (sB : set B) : set (Id U A B) =
- substInv U set (Id U A B) (equiv A B) (corrUniv A B) (rem A B sB)
+setPath (A B : U) (sB : set B) : set (Path U A B) =
+ substInv U set (Path U A B) (equiv A B) (corrUniv A B) (rem A B sB)
where
- rem (A B : U) (sB:set B) (t u:equiv A B) : prop (Id (equiv A B) t u)
- = substInv U prop (Id (equiv A B) t u) (Id (A -> B) t.1 u.1)
+ rem (A B : U) (sB:set B) (t u:equiv A B) : prop (Path (equiv A B) t u)
+ = substInv U prop (Path (equiv A B) t u) (Path (A -> B) t.1 u.1)
(eqEquivFst A B t u) (setFun A B sB t.1 u.1)
-- the collection of all sets is a groupoid
groupoidSET : groupoid SET = \(A B : SET) ->
- let rem : set (Id U A.1 B.1) = setId A.1 B.1 B.2
- rem1 : Id U (Id SET A B) (Id U A.1 B.1) =
+ let rem : set (Path U A.1 B.1) = setPath A.1 B.1 B.2
+ rem1 : Path U (Path SET A B) (Path U A.1 B.1) =
lemSigProp U set setIsProp A B
- in substInv U set (Id SET A B) (Id U A.1 B.1) rem1 rem
+ in substInv U set (Path SET A B) (Path U A.1 B.1) rem1 rem
* (ft0 : (n : nat) -> ob (suc n) -> ob n)
* (let ft (x : A) : A = mkFt ob ft0 x.1 x.2
in
- (p : (x y : A) -> Id A (ft x) y -> hom x y)
+ (p : (x y : A) -> Path A (ft x) y -> hom x y)
* (sq : (n : nat) -> (X : ob (suc n)) -> (Y : A) ->
(f : hom Y (n, ft0 n X)) ->
- (f_star : ob (suc Y.1)) * (ftf : Id A (Y.1, ft0 Y.1 f_star) Y)
+ (f_star : ob (suc Y.1)) * (ftf : Path A (Y.1, ft0 Y.1 f_star) Y)
* (q : hom (suc Y.1, f_star) (suc n, X))
- * Id (hom (suc Y.1, f_star) (n, ft0 n X))
+ * Path (hom (suc Y.1, f_star) (n, ft0 n X))
(cC C (suc Y.1, f_star) Y (n, ft0 n X) (p (suc Y.1, f_star) Y ftf) f)
(cC C (suc Y.1, f_star) (suc n, X) (n, ft0 n X) q (p (suc n, X) (n, ft0 n X) (<_> (n, ft0 n X))))
)
- * (sqId : (n : nat) -> (X : ob (suc n)) ->
- (p0 : Id A (suc n, (sq n X (n, ft0 n X) (cId C (n, ft0 n X))).1) (suc n, X))
- * (IdP (<i>cH C (p0@i) (suc n, X)) (sq n X (n, ft0 n X) (cId C (n, ft0 n X))).2.2.1 (cId C (suc n, X)))
+ * (sqPath : (n : nat) -> (X : ob (suc n)) ->
+ (p0 : Path A (suc n, (sq n X (n, ft0 n X) (cPath C (n, ft0 n X))).1) (suc n, X))
+ * (PathP (<i>cH C (p0@i) (suc n, X)) (sq n X (n, ft0 n X) (cPath C (n, ft0 n X))).2.2.1 (cPath C (suc n, X)))
)
* ((n : nat) -> (X : ob (suc n)) ->
(Y : A) -> (f : hom Y (n, ft0 n X)) ->
(Z : A) -> (g : hom Z Y) ->
(let f_star : ob (suc Y.1) = (sq n X Y f).1
g' : hom Z (Y.1, ft0 Y.1 f_star) = transport (<i>cH C Z ((sq n X Y f).2.1@-i)) g
- in (p0 : Id A (suc Z.1, (sq Y.1 f_star Z g').1)
+ in (p0 : Path A (suc Z.1, (sq Y.1 f_star Z g').1)
(suc Z.1, (sq n X Z (cC C Z Y (n, ft0 n X) g f)).1))
- * IdP (<i> cH C (p0@i) (suc n, X))
+ * PathP (<i> cH C (p0@i) (suc n, X))
(cC C (suc Z.1, (sq Y.1 f_star Z g').1) (suc Y.1, f_star) (suc n, X)
(sq Y.1 f_star Z g').2.2.1 (sq n X Y f).2.2.1)
(sq n X Z (cC C Z Y (n, ft0 n X) g f)).2.2.1)))
c0ASet (C : C0System) : set (cA (c0C C)) = setSig nat C.1 natSet C.2.2.2.1
c0Ft (C : C0System) (x : cA (c0C C)) : cA (c0C C) = mkFt C.1 C.2.2.2.2.1 x.1 x.2
-c0P (C : C0System) : (x y : cA (c0C C)) -> Id (cA (c0C C)) (c0Ft C x) y -> C.2.1 x y = C.2.2.2.2.2.1
+c0P (C : C0System) : (x y : cA (c0C C)) -> Path (cA (c0C C)) (c0Ft C x) y -> C.2.1 x y = C.2.2.2.2.2.1
c0CanSq (C : C0System) : U
= (n : nat) * (X : C.1 (suc n)) * (Y : cA (c0C C)) * (C.2.1 Y (n, C.2.2.2.2.1 n X))
c0Star (C : C0System) (T : c0CanSq C) : cA (c0C C)
= (suc T.2.2.1.1, (C.2.2.2.2.2.2.1 T.1 T.2.1 T.2.2.1 T.2.2.2).1)
c0FtStar (C : C0System) (T : c0CanSq C)
- : Id (cA (c0C C)) (c0Ft C (c0Star C T)) T.2.2.1
+ : Path (cA (c0C C)) (c0Ft C (c0Star C T)) T.2.2.1
= (C.2.2.2.2.2.2.1 T.1 T.2.1 T.2.2.1 T.2.2.2).2.1
c0Q (C : C0System) (T : c0CanSq C) : C.2.1 (c0Star C T) (suc T.1, T.2.1)
c0Square (C : C0System) (T : c0CanSq C)
: (let X : cA (c0C C) = (suc T.1, T.2.1) in
- Id (C.2.1 (c0Star C T) (c0Ft C X))
+ Path (C.2.1 (c0Star C T) (c0Ft C X))
(cC (c0C C) (c0Star C T) T.2.2.1 (c0Ft C X) (c0P C (c0Star C T) T.2.2.1 (c0FtStar C T)) T.2.2.2)
(cC (c0C C) (c0Star C T) X (c0Ft C X) (c0Q C T) (c0P C X (c0Ft C X) (<_> (c0Ft C X)))))
= (C.2.2.2.2.2.2.1 T.1 T.2.1 T.2.2.1 T.2.2.2).2.2.2
ucEquiv (A B : uc) : U
= (e : catEquiv A.1 B.1)
- * (V : Id (cA B.1) (e.1.1 A.2.2.2.1) B.2.2.2.1)
- * (VT : Id (cA B.1) (e.1.1 A.2.2.2.2.1) B.2.2.2.2.1)
- * (IdP (<i>cH B.1 (VT@i) (V@i)) (e.1.2.1 A.2.2.2.2.1 A.2.2.2.1 A.2.2.2.2.2.1) B.2.2.2.2.2.1)
+ * (V : Path (cA B.1) (e.1.1 A.2.2.2.1) B.2.2.2.1)
+ * (VT : Path (cA B.1) (e.1.1 A.2.2.2.2.1) B.2.2.2.2.1)
+ * (PathP (<i>cH B.1 (VT@i) (V@i)) (e.1.2.1 A.2.2.2.2.1 A.2.2.2.1 A.2.2.2.2.2.1) B.2.2.2.2.2.1)
-ucEquivId (A B : uc) (e : ucEquiv A B) : Id uc A B
- = let p : Id ((C:category)*catEquiv A.1 C.1) ((A.1, A.2.1), catIdEquiv A.1) ((B.1, B.2.1), e.1)
- = isContrProp ((C:category)*catEquiv A.1 C.1) (catEquivEqId' (A.1, A.2.1)) ((A.1, A.2.1), catIdEquiv A.1) ((B.1, B.2.1), e.1)
- pV : IdP (<i> cA (p@i).1.1) A.2.2.2.1 B.2.2.2.1
+ucEquivPath (A B : uc) (e : ucEquiv A B) : Path uc A B
+ = let p : Path ((C:category)*catEquiv A.1 C.1) ((A.1, A.2.1), catPathEquiv A.1) ((B.1, B.2.1), e.1)
+ = isContrProp ((C:category)*catEquiv A.1 C.1) (catEquivEqPath' (A.1, A.2.1)) ((A.1, A.2.1), catPathEquiv A.1) ((B.1, B.2.1), e.1)
+ pV : PathP (<i> cA (p@i).1.1) A.2.2.2.1 B.2.2.2.1
= <i> comp (<_> cA (p@i).1.1) ((p@i).2.1.1 A.2.2.2.1)
[(i=0)-><_>A.2.2.2.1 ,(i=1)-><k>e.2.1@k]
- pVT : IdP (<i> cA (p@i).1.1) A.2.2.2.2.1 B.2.2.2.2.1
+ pVT : PathP (<i> cA (p@i).1.1) A.2.2.2.2.1 B.2.2.2.2.1
= <i> comp (<_> cA (p@i).1.1) ((p@i).2.1.1 A.2.2.2.2.1)
[(i=0)-><_>A.2.2.2.2.1 ,(i=1)-><k>e.2.2.1@k]
- pP : IdP (<i> cH (p@i).1.1 (pVT@i) (pV@i)) A.2.2.2.2.2.1 B.2.2.2.2.2.1
+ pP : PathP (<i> cH (p@i).1.1 (pVT@i) (pV@i)) A.2.2.2.2.2.1 B.2.2.2.2.2.1
= <i> comp (<k> cH (p@i).1.1
(fill (<_> cA (p@i).1.1) ((p@i).2.1.1 A.2.2.2.2.1) [(i=0)-><_>A.2.2.2.2.1 ,(i=1)-><k>e.2.2.1@k]@k)
(fill (<_> cA (p@i).1.1) ((p@i).2.1.1 A.2.2.2.1) [(i=0)-><_>A.2.2.2.1 ,(i=1)-><k>e.2.1@k]@k))
[(i=0)-><_>A.2.2.2.2.2.1, (i=1)-><k>e.2.2.2@k]
in <i> ((p@i).1.1
,(p@i).1.2
- ,lemIdPProp (hasFinal A.1)
+ ,lemPathPProp (hasFinal A.1)
(hasFinal B.1)
(hasFinalProp A.1 A.2.1)
(<i>hasFinal (p@i).1.1)
A.2.2.1 B.2.2.1 @ i
,pV@i, pVT@i, pP@i
- ,lemIdPProp ((f : homTo A.1 (pV@0)) -> hasPullback A.1 (pV@0, f, pVT@0, pP@0))
+ ,lemPathPProp ((f : homTo A.1 (pV@0)) -> hasPullback A.1 (pV@0, f, pVT@0, pP@0))
((f : homTo B.1 (pV@1)) -> hasPullback B.1 (pV@1, f, pVT@1, pP@1))
(propPi (homTo A.1 (pV@0)) (\(f : homTo A.1 (pV@0)) -> hasPullback A.1 (pV@0, f, pVT@0, pP@0))
(\(f : homTo A.1 (pV@0)) -> hasPullbackProp A.1 A.2.1 (pV@0, f, pVT@0, pP@0)))
intD (x : obD) : cA C.1 = int x.1 x.2
homD (a b : obD) : U = C.1.1.2 (intD a) (intD b)
homDSet (a b : obD) : set (homD a b) = cHSet C.1 (intD a) (intD b)
- DId (a : obD) : homD a a = cId C.1 (intD a)
+ DPath (a : obD) : homD a a = cPath C.1 (intD a)
DC (a b c : obD) (f : homD a b) (g : homD b c) : homD a c = cC C.1 (intD a) (intD b) (intD c) f g
- DIdL (a b : obD) (g : homD a b) : Id (homD a b) (DC a a b (DId a) g) g = C.1.2.2.2.2.1 (intD a) (intD b) g
- DIdR (a b : obD) (g : homD a b) : Id (homD a b) (DC a b b g (DId b)) g = C.1.2.2.2.2.2.1 (intD a) (intD b) g
- DIdC (a b c d : obD) (f : homD a b) (g : homD b c) (h : homD c d)
- : Id (homD a d) (DC a c d (DC a b c f g) h) (DC a b d f (DC b c d g h))
+ DPathL (a b : obD) (g : homD a b) : Path (homD a b) (DC a a b (DPath a) g) g = C.1.2.2.2.2.1 (intD a) (intD b) g
+ DPathR (a b : obD) (g : homD a b) : Path (homD a b) (DC a b b g (DPath b)) g = C.1.2.2.2.2.2.1 (intD a) (intD b) g
+ DPathC (a b c d : obD) (f : homD a b) (g : homD b c) (h : homD c d)
+ : Path (homD a d) (DC a c d (DC a b c f g) h) (DC a b d f (DC b c d g h))
= C.1.2.2.2.2.2.2 (intD a) (intD b) (intD c) (intD d) f g h
DD : categoryData = (obD, homD)
D : isPrecategory DD
- = (DId, DC, homDSet, DIdL, DIdR, DIdC)
+ = (DPath, DC, homDSet, DPathL, DPathR, DPathC)
DC : precategory = (DD, D)
ft0 (n : nat) (x : ob (suc n)) : ob n = x.1
ft (x : obD) : obD = mkFt ob ft0 x.1 x.2
p0 : (n : nat) -> (x : ob n) -> homD (n, x) (ft (n, x)) = split
- zero -> \(A:Unit) -> DId (zero, A)
+ zero -> \(A:Unit) -> DPath (zero, A)
suc n -> \(X:ob (suc n)) -> (pb n X).1.2.1
- pD (x y : obD) (p : Id obD (ft x) y) : homD x y = transport (<i>homD x (p@i)) (p0 x.1 x.2)
+ pD (x y : obD) (p : Path obD (ft x) y) : homD x y = transport (<i>homD x (p@i)) (p0 x.1 x.2)
fstar (n : nat) (X : ob (suc n)) (Y : obD) (f : homD Y (n, X.1)) : obD
= (suc Y.1, Y.2, cC C.1 (intD Y) (int n X.1) V f X.2)
q_ (n : nat) (X : ob (suc n)) (Y : obD) (f : homD Y (n, X.1))
: (q : homD (fstar n X Y f) (suc n, X))
- * (qSq : Id (homD (fstar n X Y f) (n, X.1))
+ * (qSq : Path (homD (fstar n X Y f) (n, X.1))
(cC DC (fstar n X Y f) Y (n, X.1) (p0 (suc Y.1) (fstar n X Y f).2) f)
(cC DC (fstar n X Y f) (suc n, X) (n, X.1) q (p0 (suc n) X)))
- * (_ : Id (cH C.1 (intD (fstar n X Y f)) VT)
+ * (_ : Path (cH C.1 (intD (fstar n X Y f)) VT)
(cC C.1 (intD (fstar n X Y f)) (int (suc n) X) VT q (pb n X).1.2.2.1)
(pb Y.1 (Y.2, cC C.1 (intD Y) (int n X.1) V f X.2)).1.2.2.1)
* isPullback DC ((n, X.1), (Y, f), ((suc n, X), p0 (suc n) X)) (fstar n X Y f, p0 (suc Y.1) (fstar n X Y f).2, q, qSq)
gF : cH C.1 (intD Y) V = cC C.1 (intD Y) (int n X.1) V f X.2
qq : cH C.1 if_star VT = (pb Y.1 (Y.2, gF)).1.2.2.1
pg : cH C.1 if_star (int n X.1) = cC C.1 if_star (intD Y) (int n X.1) (p0 (suc Y.1) f_star.2) f
- hole0 : Id (cH C.1 if_star V)
+ hole0 : Path (cH C.1 if_star V)
(cC C.1 if_star (int n X.1) V pg X.2)
(cC C.1 if_star VT V qq p)
- = compId (cH C.1 if_star V)
+ = compPath (cH C.1 if_star V)
(cC C.1 if_star (int n X.1) V pg X.2)
(cC C.1 if_star (intD Y) V (pb Y.1 (Y.2, gF)).1.2.1 gF)
(cC C.1 if_star VT V qq p)
- (cIdC C.1 if_star (intD Y) (int n X.1) V (pb Y.1 (Y.2, gF)).1.2.1 f X.2)
+ (cPathC C.1 if_star (intD Y) (int n X.1) V (pb Y.1 (Y.2, gF)).1.2.1 f X.2)
(pb Y.1 (Y.2, gF)).1.2.2.2
pp : cospanCone C.1 (cs n X)
= (if_star, pg, qq, hole0)
q : homD (fstar n X Y f) (suc n, X) = ((pb n X).2 pp).1.1
- hole1 : Id (cH C.1 if_star V)
+ hole1 : Path (cH C.1 if_star V)
(cC C.1 if_star (intD Y) V (pb Y.1 (Y.2, gF)).1.2.1 gF)
(cC C.1 if_star VT V (cC C.1 if_star (int (suc n) X) VT q (pb n X).1.2.2.1) p)
= transport
- (<i> Id (cH C.1 if_star V)
+ (<i> Path (cH C.1 if_star V)
(cC C.1 if_star (intD Y) V (pb Y.1 (Y.2, gF)).1.2.1 gF)
(cC C.1 if_star VT V (((pb n X).2 pp).1.2.2 @ -i) p))
(pb Y.1 (Y.2, gF)).1.2.2.2
((pb Y.1 f_star.2).1.1
,(pb Y.1 f_star.2).1.2.1
,((pb n X).2 pp).1.2.2@-i
- ,lemIdPProp
- (Id (cH C.1 if_star V) (cC C.1 if_star (intD Y) V (pb Y.1 (Y.2, gF)).1.2.1 gF)
+ ,lemPathPProp
+ (Path (cH C.1 if_star V) (cC C.1 if_star (intD Y) V (pb Y.1 (Y.2, gF)).1.2.1 gF)
(cC C.1 if_star VT V (((pb n X).2 pp).1.2.2@1) p))
- (Id (cH C.1 if_star V) (cC C.1 if_star (intD Y) V (pb Y.1 (Y.2, gF)).1.2.1 gF)
+ (Path (cH C.1 if_star V) (cC C.1 if_star (intD Y) V (pb Y.1 (Y.2, gF)).1.2.1 gF)
(cC C.1 if_star VT V (((pb n X).2 pp).1.2.2@0) p))
(cHSet C.1 if_star V (cC C.1 if_star (intD Y) V (pb Y.1 (Y.2, gF)).1.2.1 gF)
(cC C.1 if_star VT V (((pb n X).2 pp).1.2.2@1) p))
- (<i>Id (cH C.1 if_star V) (cC C.1 if_star (intD Y) V (pb Y.1 (Y.2, gF)).1.2.1 gF)
+ (<i>Path (cH C.1 if_star V) (cC C.1 if_star (intD Y) V (pb Y.1 (Y.2, gF)).1.2.1 gF)
(cC C.1 if_star VT V (((pb n X).2 pp).1.2.2@-i) p))
(pb Y.1 f_star.2).1.2.2.2 hole1 @ i
))
)
sqD (n : nat) (X : ob (suc n)) (Y : obD) (f : homD Y (n, X.1))
- : (f_star : ob (suc Y.1)) * (ftf : Id obD (Y.1, ft0 Y.1 f_star) Y)
+ : (f_star : ob (suc Y.1)) * (ftf : Path obD (Y.1, ft0 Y.1 f_star) Y)
* (q : homD (suc Y.1, f_star) (suc n, X))
- * Id (homD (suc Y.1, f_star) (n, X.1))
+ * Path (homD (suc Y.1, f_star) (n, X.1))
(cC DC (suc Y.1, f_star) Y (n, X.1) (pD (suc Y.1, f_star) Y ftf) f)
(cC DC (suc Y.1, f_star) (suc n, X) (n, X.1) q (pD (suc n, X) (n, X.1) (<_> (n, X.1))))
= ((fstar n X Y f).2, <_> Y, (q_ n X Y f).1,
transport
- (<i> Id (homD (fstar n X Y f) (n, X.1))
+ (<i> Path (homD (fstar n X Y f) (n, X.1))
(cC DC (fstar n X Y f) Y (n, X.1) (transRefl (cH DC (fstar n X Y f) Y) (p0 (suc Y.1) (fstar n X Y f).2) @ -i) f)
(cC DC (fstar n X Y f) (suc n, X) (n, X.1) (q_ n X Y f).1 (transRefl (cH DC (suc n, X) (n, X.1)) (p0 (suc n) X) @ -i)))
(q_ n X Y f).2.1)
- qId (n : nat) (X : ob (suc n)) :
- (p0 : Id obD (fstar n X (n, X.1) (cId DC (n, X.1))) (suc n, X))
- * (IdP (<i>cH DC (p0@i) (suc n, X)) (q_ n X (n, X.1) (cId DC (n, X.1))).1 (cId DC (suc n, X)))
- = let f_star : obD = fstar n X (n, X.1) (cId DC (n, X.1))
- p1 : Id obD f_star (suc n, X) = <i> (suc n, X.1, cIdL C.1 (int n X.1) V X.2 @ i)
+ qPath (n : nat) (X : ob (suc n)) :
+ (p0 : Path obD (fstar n X (n, X.1) (cPath DC (n, X.1))) (suc n, X))
+ * (PathP (<i>cH DC (p0@i) (suc n, X)) (q_ n X (n, X.1) (cPath DC (n, X.1))).1 (cPath DC (suc n, X)))
+ = let f_star : obD = fstar n X (n, X.1) (cPath DC (n, X.1))
+ p1 : Path obD f_star (suc n, X) = <i> (suc n, X.1, cPathL C.1 (int n X.1) V X.2 @ i)
if_star : cA C.1 = intD f_star
- gF : cH C.1 (int n X.1) V = cC C.1 (int n X.1) (int n X.1) V (cId DC (n, X.1)) X.2
+ gF : cH C.1 (int n X.1) V = cC C.1 (int n X.1) (int n X.1) V (cPath DC (n, X.1)) X.2
qq : cH C.1 if_star VT = (pb n (X.1, gF)).1.2.2.1
pg : cH C.1 if_star (int n X.1)
- = cC C.1 if_star (int n X.1) (int n X.1) (p0 (suc n) f_star.2) (cId DC (n, X.1))
- hole0 : Id (cH C.1 if_star V)
+ = cC C.1 if_star (int n X.1) (int n X.1) (p0 (suc n) f_star.2) (cPath DC (n, X.1))
+ hole0 : Path (cH C.1 if_star V)
(cC C.1 if_star (int n X.1) V pg X.2)
(cC C.1 if_star VT V qq p)
- = compId (cH C.1 if_star V)
+ = compPath (cH C.1 if_star V)
(cC C.1 if_star (int n X.1) V pg X.2)
(cC C.1 if_star (int n X.1) V (pb n (X.1, gF)).1.2.1 gF)
(cC C.1 if_star VT V qq p)
- (cIdC C.1 if_star (int n X.1) (int n X.1) V (pb n (X.1, gF)).1.2.1 (cId DC (n, X.1)) X.2)
+ (cPathC C.1 if_star (int n X.1) (int n X.1) V (pb n (X.1, gF)).1.2.1 (cPath DC (n, X.1)) X.2)
(pb n (X.1, gF)).1.2.2.2
pp : cospanCone C.1 (cs n X)
= (if_star, pg, qq, hole0)
- ppId : Id (cospanCone C.1 (cs n X)) pp (pb n X).1
+ ppPath : Path (cospanCone C.1 (cs n X)) pp (pb n X).1
= <i> (intD (p1@i),
- cIdR C.1 (intD (p1@i)) (int n X.1) (p0 (suc n) (p1@i).2) @ i,
- (pb n (X.1, cIdL C.1 (int n X.1) V X.2 @ i)).1.2.2.1,
- lemIdPProp (Id (cH C.1 if_star V) (cC C.1 if_star (int n X.1) V pg X.2) (cC C.1 if_star VT V qq p))
- (Id (cH C.1 (int (suc n) X) V)
+ cPathR C.1 (intD (p1@i)) (int n X.1) (p0 (suc n) (p1@i).2) @ i,
+ (pb n (X.1, cPathL C.1 (int n X.1) V X.2 @ i)).1.2.2.1,
+ lemPathPProp (Path (cH C.1 if_star V) (cC C.1 if_star (int n X.1) V pg X.2) (cC C.1 if_star VT V qq p))
+ (Path (cH C.1 (int (suc n) X) V)
(cC C.1 (int (suc n) X) (int n X.1) V (p0 (suc n) X) X.2)
(cC C.1 (int (suc n) X) VT V (pb n X).1.2.2.1 p))
(cHSet C.1 if_star V (cC C.1 if_star (int n X.1) V pg X.2) (cC C.1 if_star VT V qq p))
- (<i>Id (cH C.1 (intD (p1@i)) V)
- (cC C.1 (intD (p1@i)) (int n X.1) V (cIdR C.1 (intD (p1@i)) (int n X.1) (p0 (suc n) (p1@i).2) @ i) X.2)
- (cC C.1 (intD (p1@i)) VT V ((pb n (X.1, cIdL C.1 (int n X.1) V X.2 @ i)).1.2.2.1) p))
+ (<i>Path (cH C.1 (intD (p1@i)) V)
+ (cC C.1 (intD (p1@i)) (int n X.1) V (cPathR C.1 (intD (p1@i)) (int n X.1) (p0 (suc n) (p1@i).2) @ i) X.2)
+ (cC C.1 (intD (p1@i)) VT V ((pb n (X.1, cPathL C.1 (int n X.1) V X.2 @ i)).1.2.2.1) p))
hole0
(pb n X).1.2.2.2
@ i
)
pph : cospanConeHom C.1 (cs n X) pp (pb n X).1
- = transport (<i> cospanConeHom C.1 (cs n X) (ppId@-i) (pb n X).1) (cospanConeId C.1 (cs n X) (pb n X).1)
- pphId : Id (cospanConeHom C.1 (cs n X) pp (pb n X).1) ((pb n X).2 pp).1 pph
+ = transport (<i> cospanConeHom C.1 (cs n X) (ppPath@-i) (pb n X).1) (cospanConePath C.1 (cs n X) (pb n X).1)
+ pphPath : Path (cospanConeHom C.1 (cs n X) pp (pb n X).1) ((pb n X).2 pp).1 pph
= ((pb n X).2 pp).2 pph
- qId : Id (cH DC f_star (suc n, X)) ((pb n X).2 pp).1.1 (transport (<i> cH C.1 (ppId@-i).1 (int (suc n) X)) (cId DC (suc n, X)))
- = <i>(pphId@i).1
+ qPath : Path (cH DC f_star (suc n, X)) ((pb n X).2 pp).1.1 (transport (<i> cH C.1 (ppPath@-i).1 (int (suc n) X)) (cPath DC (suc n, X)))
+ = <i>(pphPath@i).1
in
( p1
- , <i> substIdP
+ , <i> substPathP
(cH DC (p1@1) (suc n, X)) (cH DC (p1@0) (suc n, X))
- (<i>cH DC (p1@-i) (suc n, X)) (cId DC (suc n, X)) (q_ n X (n, X.1) (cId DC (n, X.1))).1
- (<i>qId@-i)
+ (<i>cH DC (p1@-i) (suc n, X)) (cPath DC (suc n, X)) (q_ n X (n, X.1) (cPath DC (n, X.1))).1
+ (<i>qPath@-i)
@ -i
)
qComp (n : nat) (X : ob (suc n))
(Y : obD) (f : homD Y (n, X.1))
(Z : obD) (g : homD Z Y)
- : (p0 : Id obD (fstar Y.1 (fstar n X Y f).2 Z g)
+ : (p0 : Path obD (fstar Y.1 (fstar n X Y f).2 Z g)
(fstar n X Z (cC DC Z Y (n, X.1) g f)))
- * IdP (<i> cH DC (p0@i) (suc n, X))
+ * PathP (<i> cH DC (p0@i) (suc n, X))
(cC DC (fstar Y.1 (fstar n X Y f).2 Z g) (fstar n X Y f) (suc n, X)
(q_ Y.1 (fstar n X Y f).2 Z g).1 (q_ n X Y f).1)
(q_ n X Z (cC DC Z Y (n, X.1) g f)).1
gF : cH C.1 (intD Z) V = cC C.1 (intD Z) (int n X.1) V F X.2
qq : cH C.1 if_star VT = (pb Z.1 (Z.2, gF)).1.2.2.1
pg : cH C.1 if_star (int n X.1) = cC C.1 if_star (intD Z) (int n X.1) (p0 (suc Z.1) f_star.2) F
- hole0 : Id (cH C.1 if_star V)
+ hole0 : Path (cH C.1 if_star V)
(cC C.1 if_star (int n X.1) V pg X.2)
(cC C.1 if_star VT V qq p)
- = compId (cH C.1 if_star V)
+ = compPath (cH C.1 if_star V)
(cC C.1 if_star (int n X.1) V pg X.2)
(cC C.1 if_star (intD Z) V (pb Z.1 (Z.2, gF)).1.2.1 gF)
(cC C.1 if_star VT V qq p)
- (cIdC C.1 if_star (intD Z) (int n X.1) V (pb Z.1 (Z.2, gF)).1.2.1 F X.2)
+ (cPathC C.1 if_star (intD Z) (int n X.1) V (pb Z.1 (Z.2, gF)).1.2.1 F X.2)
(pb Z.1 (Z.2, gF)).1.2.2.2
pp : cospanCone C.1 (cs n X)
= (if_star, pg, qq, hole0)
= cC DC f_star2 Z (n,X.1) (p0 (suc Z.1) f_star2.2) F
qq2 : cH C.1 if_star2 VT
= (pb Z.1 (Z.2, cC C.1 (intD Z) (intD Y) V g (cC C.1 (intD Y) (int n X.1) V f X.2))).1.2.2.1
- p1 : Id obD f_star2 f_star
- = <i> (suc Z.1, Z.2, cIdC C.1 (intD Z) (intD Y) (int n X.1) V g f X.2 @ -i)
+ p1 : Path obD f_star2 f_star
+ = <i> (suc Z.1, Z.2, cPathC C.1 (intD Z) (intD Y) (int n X.1) V g f X.2 @ -i)
pp2 : cospanCone C.1 (cs n X)
= (if_star2, pg2, qq2,
- transport (<i>Id (cH C.1 (intD(p1@-i)) V)
+ transport (<i>Path (cH C.1 (intD(p1@-i)) V)
(cC C.1 (intD(p1@-i)) (int n X.1) V (cC DC (p1@-i) Z (n,X.1) (p0 (suc Z.1) (p1@-i).2) F) X.2)
- (cC C.1 (intD(p1@-i)) VT V (pb Z.1 (Z.2, cIdC C.1 (intD Z) (intD Y) (int n X.1) V g f X.2 @ i)).1.2.2.1 p))
+ (cC C.1 (intD(p1@-i)) VT V (pb Z.1 (Z.2, cPathC C.1 (intD Z) (intD Y) (int n X.1) V g f X.2 @ i)).1.2.2.1 p))
hole0)
- ppId : Id (cospanCone C.1 (cs n X)) pp2 pp
+ ppPath : Path (cospanCone C.1 (cs n X)) pp2 pp
= <i> (intD (p1@i),
cC DC (p1@i) Z (n,X.1) (p0 (suc Z.1) (p1@i).2) F,
- (pb Z.1 (Z.2, cIdC C.1 (intD Z) (intD Y) (int n X.1) V g f X.2 @ -i)).1.2.2.1,
- lemIdPProp
- (Id (cH C.1 if_star2 V) (cC C.1 if_star2 (int n X.1) V pg2 X.2) (cC C.1 if_star2 VT V qq2 p))
- (Id (cH C.1 if_star V) (cC C.1 if_star (int n X.1) V pg X.2) (cC C.1 if_star VT V qq p))
+ (pb Z.1 (Z.2, cPathC C.1 (intD Z) (intD Y) (int n X.1) V g f X.2 @ -i)).1.2.2.1,
+ lemPathPProp
+ (Path (cH C.1 if_star2 V) (cC C.1 if_star2 (int n X.1) V pg2 X.2) (cC C.1 if_star2 VT V qq2 p))
+ (Path (cH C.1 if_star V) (cC C.1 if_star (int n X.1) V pg X.2) (cC C.1 if_star VT V qq p))
(cHSet C.1 if_star2 V (cC C.1 if_star2 (int n X.1) V pg2 X.2) (cC C.1 if_star2 VT V qq2 p))
- (<i>Id (cH C.1 (intD(p1@i)) V)
+ (<i>Path (cH C.1 (intD(p1@i)) V)
(cC C.1 (intD(p1@i)) (int n X.1) V (cC DC (p1@i) Z (n,X.1) (p0 (suc Z.1) (p1@i).2) F) X.2)
- (cC C.1 (intD(p1@i)) VT V (pb Z.1 (Z.2, cIdC C.1 (intD Z) (intD Y) (int n X.1) V g f X.2 @ -i)).1.2.2.1 p))
+ (cC C.1 (intD(p1@i)) VT V (pb Z.1 (Z.2, cPathC C.1 (intD Z) (intD Y) (int n X.1) V g f X.2 @ -i)).1.2.2.1 p))
pp2.2.2.2 hole0 @ i)
- hole3 : Id (cH C.1 if_star2 (int n X.1))
+ hole3 : Path (cH C.1 if_star2 (int n X.1))
(cC DC f_star2 (suc n, X) (n,X.1)
(cC C.1 if_star2 (intD (fstar n X Y f)) (int (suc n) X) q2 (q_ n X Y f).1)
(p0 (suc n) X)
)
(cC DC f_star2 Z (n,X.1) (p0 (suc Z.1) f_star2.2) F)
= undefined
- -- = compId (cH C.1 if_star2 (int n X.1))
+ -- = compPath (cH C.1 if_star2 (int n X.1))
-- (cC DC f_star2 (suc n, X) (n,X.1)
-- (cC DC f_star2 (fstar n X Y f) (suc n, X) q2 (q_ n X Y f).1)
-- (p0 (suc n) X))
-- q2
-- (cC DC (fstar n X Y f) (suc n, X) (n, X.1) (q_ n X Y f).1 (p0 (suc n) X)))
-- (cC DC f_star2 Z (n,X.1) (p0 (suc Z.1) f_star2.2) F)
- -- (cIdC DC f_star2 (fstar n X Y f) (suc n, X) (n,X.1) q2 (q_ n X Y f).1 (p0 (suc n) X))
- -- (compId (cH C.1 if_star2 (int n X.1))
+ -- (cPathC DC f_star2 (fstar n X Y f) (suc n, X) (n,X.1) q2 (q_ n X Y f).1 (p0 (suc n) X))
+ -- (compPath (cH C.1 if_star2 (int n X.1))
-- (cC DC f_star2 (fstar n X Y f) (n,X.1)
-- q2
-- (cC DC (fstar n X Y f) (suc n, X) (n, X.1) (q_ n X Y f).1 (p0 (suc n) X)))
-- (cC DC f_star2 Z (n,X.1) (p0 (suc Z.1) f_star2.2) F)
-- (<i> cC DC f_star2 (fstar n X Y f) (n,X.1) q2
-- ((q_ n X Y f).2.1@-i))
- -- (compId (cH C.1 if_star2 (int n X.1))
+ -- (compPath (cH C.1 if_star2 (int n X.1))
-- (cC DC f_star2 (fstar n X Y f) (n,X.1)
-- q2
-- (cC DC (fstar n X Y f) Y (n, X.1) (p0 (suc Y.1) (fstar n X Y f).2) f))
-- (cC DC f_star2 (fstar n X Y f) Y q2 (p0 (suc Y.1) (fstar n X Y f).2))
-- f)
-- (cC DC f_star2 Z (n,X.1) (p0 (suc Z.1) f_star2.2) F)
- -- (<i>cIdC DC f_star2 (fstar n X Y f) Y (n, X.1) q2 (p0 (suc Y.1) (fstar n X Y f).2) f@-i)
- -- (compId (cH C.1 if_star2 (int n X.1))
+ -- (<i>cPathC DC f_star2 (fstar n X Y f) Y (n, X.1) q2 (p0 (suc Y.1) (fstar n X Y f).2) f@-i)
+ -- (compPath (cH C.1 if_star2 (int n X.1))
-- (cC DC f_star2 Y (n,X.1)
-- (cC DC f_star2 (fstar n X Y f) Y q2 (p0 (suc Y.1) (fstar n X Y f).2))
-- f)
-- (cC DC f_star2 Z (n,X.1) (p0 (suc Z.1) f_star2.2) F)
-- (<i>cC DC f_star2 Y (n,X.1)
-- ((q_ Y.1 (fstar n X Y f).2 Z g).2.1@-i) f)
- -- (cIdC DC f_star2 Z Y (n, X.1) (p0 (suc Z.1) f_star2.2) g f))))
+ -- (cPathC DC f_star2 Z Y (n, X.1) (p0 (suc Z.1) f_star2.2) g f))))
-- opaque hole3
- hole4 : Id (cH C.1 if_star2 VT)
+ hole4 : Path (cH C.1 if_star2 VT)
(cC C.1 if_star2 (int (suc n) X) VT
(cC C.1 if_star2 (intD (fstar n X Y f)) (int (suc n) X) q2 (q_ n X Y f).1)
(pb n X).1.2.2.1)
(pb Z.1 (Z.2, cC C.1 (intD Z) (intD Y) V g (cC C.1 (intD Y) (int n X.1) V f X.2))).1.2.2.1
= undefined
- -- = compId (cH C.1 if_star2 VT)
+ -- = compPath (cH C.1 if_star2 VT)
-- (cC C.1 if_star2 (int (suc n) X) VT
-- (cC C.1 if_star2 (intD (fstar n X Y f)) (int (suc n) X) q2 (q_ n X Y f).1)
-- (pb n X).1.2.2.1)
-- q2
-- (cC C.1 (intD (fstar n X Y f)) (int (suc n) X) VT (q_ n X Y f).1 (pb n X).1.2.2.1))
-- (pb Z.1 (Z.2, cC C.1 (intD Z) (intD Y) V g (cC C.1 (intD Y) (int n X.1) V f X.2))).1.2.2.1
- -- (cIdC C.1 if_star2 (intD (fstar n X Y f)) (int (suc n) X) VT q2 (q_ n X Y f).1 (pb n X).1.2.2.1)
- -- (compId (cH C.1 if_star2 VT)
+ -- (cPathC C.1 if_star2 (intD (fstar n X Y f)) (int (suc n) X) VT q2 (q_ n X Y f).1 (pb n X).1.2.2.1)
+ -- (compPath (cH C.1 if_star2 VT)
-- (cC C.1 if_star2 (intD (fstar n X Y f)) VT
-- q2
-- (cC C.1 (intD (fstar n X Y f)) (int (suc n) X) VT (q_ n X Y f).1 (pb n X).1.2.2.1))
pph : cospanConeHom C.1 (cs n X) pp (pb n X).1
= transport
- (<i>cospanConeHom C.1 (cs n X) (ppId@i) (pb n X).1)
+ (<i>cospanConeHom C.1 (cs n X) (ppPath@i) (pb n X).1)
( cC C.1 if_star2 (intD (fstar n X Y f)) (int (suc n) X) q2 (q_ n X Y f).1
, hole3
, hole4
)
- pphId : Id (cospanConeHom C.1 (cs n X) pp (pb n X).1) ((pb n X).2 pp).1 pph
+ pphPath : Path (cospanConeHom C.1 (cs n X) pp (pb n X).1) ((pb n X).2 pp).1 pph
= ((pb n X).2 pp).2 pph
- qId : Id (cH C.1 if_star (int (suc n) X))
+ qPath : Path (cH C.1 if_star (int (suc n) X))
(transport (<i>cH C.1 (intD(p1@i)) (int (suc n) X))
(cC C.1 if_star2 (intD (fstar n X Y f)) (int (suc n) X) q2 (q_ n X Y f).1))
(q_ n X Z (cC DC Z Y (n, X.1) g f)).1
- = <i>(pphId@-i).1
+ = <i>(pphPath@-i).1
in (p1,
- substIdP (cH DC (p1@0) (suc n, X)) (cH DC (p1@1) (suc n, X))
+ substPathP (cH DC (p1@0) (suc n, X)) (cH DC (p1@1) (suc n, X))
(<i> cH DC (p1@i) (suc n, X))
(cC C.1 if_star2 (intD (fstar n X Y f)) (int (suc n) X) q2 (q_ n X Y f).1)
(q_ n X Z (cC DC Z Y (n, X.1) g f)).1
- qId)
+ qPath)
qComp' (n : nat) (X : ob (suc n))
(Y : obD) (f : homD Y (n, X.1))
(Z : obD) (g : homD Z Y)
: (let g' : homD Z Y = transport (<_> homD Z Y) g in
- (p0 : Id obD (fstar Y.1 (fstar n X Y f).2 Z g')
+ (p0 : Path obD (fstar Y.1 (fstar n X Y f).2 Z g')
(fstar n X Z (cC DC Z Y (n, X.1) g f)))
- * IdP (<i> cH DC (p0@i) (suc n, X))
+ * PathP (<i> cH DC (p0@i) (suc n, X))
(cC DC (fstar Y.1 (fstar n X Y f).2 Z g') (fstar n X Y f) (suc n, X)
(q_ Y.1 (fstar n X Y f).2 Z g').1 (q_ n X Y f).1)
(q_ n X Z (cC DC Z Y (n, X.1) g f)).1)
= transport
- (<j> (p0 : Id obD (fstar Y.1 (fstar n X Y f).2 Z (transRefl (homD Z Y) g @ -j))
+ (<j> (p0 : Path obD (fstar Y.1 (fstar n X Y f).2 Z (transRefl (homD Z Y) g @ -j))
(fstar n X Z (cC DC Z Y (n, X.1) g f)))
- * IdP (<i> cH DC (p0@i) (suc n, X))
+ * PathP (<i> cH DC (p0@i) (suc n, X))
(cC DC (fstar Y.1 (fstar n X Y f).2 Z (transRefl (homD Z Y) g @ -j)) (fstar n X Y f) (suc n, X)
(q_ Y.1 (fstar n X Y f).2 Z (transRefl (homD Z Y) g @ -j)).1 (q_ n X Y f).1)
(q_ n X Z (cC DC Z Y (n, X.1) g f)).1)
(qComp n X Y f Z g)
hole : C0System
- = (ob, homD, D, obSet, ft0, pD, sqD, qId, qComp')
+ = (ob, homD, D, obSet, ft0, pD, sqD, qPath, qComp')
|- ID A B Type
-IdP is heterogeneous equality:
+PathP is heterogeneous equality:
|- P : ID A B |- a : A |- b : B
-----------------------------------------------
- |- IdP P a b Type
+ |- PathP P a b Type
-}
--- The usual identity can be seen a special case of IdP:
-Id (A : U) (a b : A) : U = IdP (<i> A) a b
+-- The usual identity can be seen a special case of PathP:
+Path (A : U) (a b : A) : U = PathP (<i> A) a b
-- "<i>" abstracts the name/color/dimension i:
-refl (A : U) (a : A) : Id A a a = <i> a
+refl (A : U) (a : A) : Path A a a = <i> a
{-
-A proof "p : Id A a b" is thought of as a line between a and b:
+A proof "p : Path A a b" is thought of as a line between a and b:
p
a ------> b
-A proof "sq : Id (Id A a b) p q" is thought of as a square between p and q:
+A proof "sq : Path (Path A a b) p q" is thought of as a square between p and q:
q
-- It is possible to take the face of a path to obtain its endpoints:
-face0 (A : U) (a b : A) (p : Id A a b) : Id A a a = refl A (p @ 0)
-face1 (A : U) (a b : A) (p : Id A a b) : Id A b b = refl A (p @ 1)
+face0 (A : U) (a b : A) (p : Path A a b) : Path A a a = refl A (p @ 0)
+face1 (A : U) (a b : A) (p : Path A a b) : Path A b b = refl A (p @ 1)
-- By applying a path to a name i, "p @ i", it is seen as a path "in
-- dimension i" connecting "p @ 0" to "p @ 1". This way we get a
-- simple proof of cong:
-cong (A B : U) (f : A -> B) (a b : A) (p : Id A a b) : Id B (f a) (f b) =
+cong (A B : U) (f : A -> B) (a b : A) (p : Path A a b) : Path B (f a) (f b) =
<i> f (p @ i)
-- This also gives a short proof of function extensionality:
funExt (A : U) (B : A -> U) (f g : (x : A) -> B x)
- (p : (x : A) -> Id (B x) (f x) (g x)) : Id ((y : A) -> B y) f g =
+ (p : (x : A) -> Path (B x) (f x) (g x)) : Path ((y : A) -> B y) f g =
<i> \(a : A) -> (p a) @ i
{-
-}
-- Applying a path to a negated name inverts it:
-sym (A : U) (a b : A) (p : Id A a b) : Id A b a = <i> p @ -i
+sym (A : U) (a b : A) (p : Path A a b) : Path A b a = <i> p @ -i
-- This operation is an involution:
-symK (A : U) (a b : A) (p : Id A a b) : Id (Id A a b) p p =
- refl (Id A a b) (sym A b a (sym A a b p))
+symK (A : U) (a b : A) (p : Path A a b) : Path (Path A a b) p p =
+ refl (Path A a b) (sym A b a (sym A a b p))
{- Connections:
-- This gives a simple proof that singletons are contractible:
-singl (A : U) (a : A) : U = (x : A) * Id A a x
+singl (A : U) (a : A) : U = (x : A) * Path A a x
-contrSingl (A : U) (a b : A) (p : Id A a b) :
- Id (singl A a) (a,refl A a) (b,p) =
+contrSingl (A : U) (a b : A) (p : Path A a b) :
+ Path (singl A a) (a,refl A a) (b,p) =
<i> (p @ i, <j> p @ i /\ j)
Note that formulas does not form a boolean algebra, but a de Morgan
algebra. This means that "p @ 0" and "p @ i /\ -i" are different!
-testDeMorganAlgebra (A : U) (a b : A) (p : Id A a b) :
- Id (Id A a a) (<_> p @ 0) (<_> p @ 0) =
- refl (Id A a a) (<i> p @ i /\ -i)
+testDeMorganAlgebra (A : U) (a b : A) (p : Path A a b) :
+ Path (Path A a a) (<_> p @ 0) (<_> p @ 0) =
+ refl (Path A a a) (<i> p @ i /\ -i)
This is clear with the intuition that /\ correspond to min. More about
this later...
-}
-testConj (A : U) (a b : A) (p : Id A a b) : Id A a a = <i> p @ i /\ -i
-testDisj (A : U) (a b : A) (p : Id A a b) : Id A b b = <i> p @ i \/ -i
+testConj (A : U) (a b : A) (p : Path A a b) : Path A a a = <i> p @ i /\ -i
+testDisj (A : U) (a b : A) (p : Path A a b) : Path A b b = <i> p @ i \/ -i
{-
side (opposite of the principal side).
-Transitivity of Id is obtained from a composition of this open square:
+Transitivity of Path is obtained from a composition of this open square:
a - - - - - - - - > c
The composition computes the dashed line at the top of the square.
-}
-trans (A : U) (a b c : A) (p : Id A a b) (q : Id A b c) : Id A a c =
+trans (A : U) (a b c : A) (p : Path A a b) (q : Path A b c) : Path A a c =
<i> comp (<_> A) (p @ i) [ (i = 0) -> <j> a, (i = 1) -> q ]
-- "Kan composition":
-kan (A : U) (a b c d : A) (p : Id A a b)
- (q : Id A a c) (r : Id A b d) : Id A c d =
+kan (A : U) (a b c d : A) (p : Path A a b)
+ (q : Path A a c) (r : Path A b d) : Path A c d =
<i> comp (<_> A) (p @ i) [ (i = 0) -> q, (i = 1) -> r ]
-- The first two nonzero h-levels are propositions and sets:
-prop (A : U) : U = (a b : A) -> Id A a b
-set (A : U) : U = (a b : A) -> prop (Id A a b)
+prop (A : U) : U = (a b : A) -> Path A a b
+set (A : U) : U = (a b : A) -> prop (Path A a b)
-- Using compositions we get a short proof that any prop is a set. To
-- understand this proof one should draw an open cube with the back
-- face being a constant square in a and the sides given by the system
-- below.
propSet (A : U) (h : prop A) : set A =
- \(a b : A) (p q : Id A a b) ->
+ \(a b : A) (p q : Path A a b) ->
<i j> comp (<_> A) a [ (j=0) -> h a a
, (j=1) -> h a b
, (i=0) -> h a (p @ j)
v
-}
-Square (A : U) (a0 a1 b0 b1 : A) (u : Id A a0 a1) (v : Id A b0 b1)
- (r0 : Id A a0 b0) (r1 : Id A a1 b1) : U
- = IdP (<i> (IdP (<j> A) (u @ i) (v @ i))) r0 r1
+Square (A : U) (a0 a1 b0 b1 : A) (u : Path A a0 a1) (v : Path A b0 b1)
+ (r0 : Path A a0 b0) (r1 : Path A a1 b1) : U
+ = PathP (<i> (PathP (<j> A) (u @ i) (v @ i))) r0 r1
-constSquare (A : U) (a : A) (p : Id A a a) : Square A a a a a p p p p =
+constSquare (A : U) (a : A) (p : Path A a a) : Square A a a a a p p p p =
<i j> comp (<_> A) a [ (i = 0) -> <k> p @ j \/ - k
, (i = 1) -> <k> p @ j /\ k
, (j = 0) -> <k> p @ i \/ - k
zero -> a
suc b -> suc (add a b)
-addZero : (a : nat) -> Id nat (add zero a) a = split
+addZero : (a : nat) -> Path nat (add zero a) a = split
zero -> <i> zero
suc a' -> <i> suc (addZero a' @ i)
-addSuc (a : nat) : (b : nat) -> Id nat (add (suc a) b) (suc (add a b)) = split
+addSuc (a : nat) : (b : nat) -> Path nat (add (suc a) b) (suc (add a b)) = split
zero -> <i> suc a
suc b' -> <i> suc (addSuc a b' @ i)
-addCom (a : nat) : (b : nat) -> Id nat (add a b) (add b a) = split
+addCom (a : nat) : (b : nat) -> Path nat (add a b) (add b a) = split
zero -> <i> addZero a @ -i
suc b' -> <i> comp (<_> nat) (suc (addCom a b' @ i)) [ (i = 0) -> <j> suc (add a b')
, (i = 1) -> <j> addSuc b' a @ -j ]
-addAssoc (a b : nat) : (c : nat) -> Id nat (add a (add b c)) (add (add a b) c) = split
+addAssoc (a b : nat) : (c : nat) -> Path nat (add a (add b c)) (add (add a b) c) = split
zero -> <i> add a b
suc c' -> <i> suc (addAssoc a b c' @ i)
Transport takes a path in the universe between A and B and produces a
function from A to B:
- transport : Id U A B -> A -> B
+ transport : Path U A B -> A -> B
It is defined internally as a composition with an empty system.
-}
-- This gives a simple proof of subst:
-subst (A : U) (P : A -> U) (a b : A) (p : Id A a b) (e : P a) : P b =
+subst (A : U) (P : A -> U) (a b : A) (p : Path A a b) (e : P a) : P b =
transport (cong A U P a b p) e
-- Combined with the fact that singletons are contractible this gives
-- the J eliminator:
-J (A : U) (a : A) (C : (x : A) -> Id A a x -> U)
- (d : C a (refl A a)) (x : A) (p : Id A a x) : C x p =
+J (A : U) (a : A) (C : (x : A) -> Path A a x -> U)
+ (d : C a (refl A a)) (x : A) (p : Path A a x) : C x p =
subst (singl A a) T (a,refl A a) (x,p) (contrSingl A a x p) d
where T (bp : singl A a) : U = C bp.1 bp.2
-- Note: Transporting with refl is not the identity function:
--- transRefl (A : U) (a : A) : Id A a a = refl A (transport (refl U A) a)
+-- transRefl (A : U) (a : A) : Path A a a = refl A (transport (refl U A) a)
-- This implies that the elimination rule for J does not hold definitonally:
--- defEqJ (A : U) (a : A) (C : (x : A) -> Id A a x -> U) (d : C a (refl A a)) :
--- Id (C a (refl A a)) (J A a C d a (refl A a)) d = refl (C a (refl A a)) d
+-- defEqJ (A : U) (a : A) (C : (x : A) -> Path A a x -> U) (d : C a (refl A a)) :
+-- Path (C a (refl A a)) (J A a C d a (refl A a)) d = refl (C a (refl A a)) d
-- This might not be too bad though. Nils-Anders Danielsson has
-- verified in Agda that many of the consequences of univalence can be
\r
import prelude\r
\r
-inlNotinr (A B:U) (a:A) (b:B) (h: Id (or A B) (inl a) (inr b)) : N0 = \r
+inlNotinr (A B:U) (a:A) (b:B) (h: Path (or A B) (inl a) (inr b)) : N0 = \r
subst (or A B) T (inl a) (inr b) h tt\r
where\r
T : or A B -> U = split\r
inl _ -> Unit\r
inr _ -> N0\r
\r
-inrNotinl (A B:U) (a:A) (b:B) (h : Id (or A B) (inr b) (inl a)) : N0 = \r
+inrNotinl (A B:U) (a:A) (b:B) (h : Path (or A B) (inr b) (inl a)) : N0 = \r
subst (or A B) T (inr b) (inl a) h tt\r
where\r
T : or A B -> U = split\r
inl _ -> N0\r
inr _ -> Unit\r
\r
-injInl (A B :U) (x0 x1:A) (h : Id (or A B) (inl x0) (inl x1)) : Id A x0 x1 = \r
+injInl (A B :U) (x0 x1:A) (h : Path (or A B) (inl x0) (inl x1)) : Path A x0 x1 = \r
subst (or A B) T (inl x0) (inl x1) h (refl A x0)\r
where\r
T : or A B -> U = split\r
- inl x -> Id A x0 x\r
+ inl x -> Path A x0 x\r
inr _ -> N0\r
\r
-injInr (A B :U) (x0 x1:B) (h: Id (or A B) (inr x0) (inr x1)) : Id B x0 x1 = \r
+injInr (A B :U) (x0 x1:B) (h: Path (or A B) (inr x0) (inr x1)) : Path B x0 x1 = \r
subst (or A B) T (inr x0) (inr x1) h (refl B x0)\r
where\r
T : or A B -> U = split\r
inl _ -> N0\r
- inr x -> Id B x0 x\r
+ inr x -> Path B x0 x\r
\r
-- If A and B are discrete then "A or B" is discrete\r
orDisc (A B : U) (dA : discrete A) (dB : discrete B) :\r
- (z z1 : or A B) -> dec (Id (or A B) z z1) = split\r
+ (z z1 : or A B) -> dec (Path (or A B) z z1) = split\r
inl a -> rem1\r
- where rem1 : (z1:or A B) -> dec (Id (or A B) (inl a) z1) = split\r
+ where rem1 : (z1:or A B) -> dec (Path (or A B) (inl a) z1) = split\r
inl a1 -> rem (dA a a1)\r
- where rem : dec (Id A a a1) -> dec (Id (or A B) (inl a) (inl a1)) = split\r
+ where rem : dec (Path A a a1) -> dec (Path (or A B) (inl a) (inl a1)) = split\r
inl p -> inl (<i> inl (p @ i))\r
- inr h -> inr (\ (p:Id (or A B) (inl a) (inl a1)) -> h (injInl A B a a1 p))\r
+ inr h -> inr (\ (p:Path (or A B) (inl a) (inl a1)) -> h (injInl A B a a1 p))\r
inr b -> inr (inlNotinr A B a b)\r
inr b -> rem1\r
- where rem1 : (z1:or A B) -> dec (Id (or A B) (inr b) z1) = split\r
+ where rem1 : (z1:or A B) -> dec (Path (or A B) (inr b) z1) = split\r
inl a -> inr (inrNotinl A B a b)\r
inr b1 -> rem (dB b b1)\r
- where rem : dec (Id B b b1) -> dec (Id (or A B) (inr b) (inr b1)) = split\r
+ where rem : dec (Path B b b1) -> dec (Path (or A B) (inr b) (inr b1)) = split\r
inl p -> inl (<i> inr (p @ i))\r
- inr h -> inr (\ (p:Id (or A B) (inr b) (inr b1)) -> h (injInr A B b b1 p))\r
+ inr h -> inr (\ (p:Path (or A B) (inr b) (inr b1)) -> h (injInr A B b b1 p))\r
\r
import prelude
fiber (A B : U) (f : A -> B) (y : B) : U =
- (x : A) * Id B y (f x)
+ (x : A) * Path B y (f x)
isEquiv (A B : U) (f : A -> B) : U = (y : B) -> isContr (fiber A B f y)
<i> \ (y:B) -> propIsContr (fiber A B f y) (u0 y) (u1 y) @ i
equivLemma (A B : U)
- : (v w : equiv A B) -> Id (A -> B) v.1 w.1 -> Id (equiv A B) v w
+ : (v w : equiv A B) -> Path (A -> B) v.1 w.1 -> Path (equiv A B) v w
= lemSig (A -> B) (isEquiv A B) (propIsEquiv A B)
idIsEquiv (A : U) : isEquiv A A (idfun A) =
\(a : A) -> ((a, refl A a)
,\(z : fiber A A (idfun A) a) -> contrSingl A a z.1 z.2)
-equivId (T A : U) (f : T -> A) (p : isEquiv T A f) : Id U T A =
+equivPath (T A : U) (f : T -> A) (p : isEquiv T A f) : Path U T A =
<i> glue A [ (i=0) -> (T,f,p), (i=1) -> (A,idfun A, idIsEquiv A)]
-- for univalence
invEq (A B:U) (w:equiv A B) (y:B) : A = (w.2 y).1.1
-retEq (A B:U) (w:equiv A B) (y:B) : Id B (w.1 (invEq A B w y)) y =
+retEq (A B:U) (w:equiv A B) (y:B) : Path B (w.1 (invEq A B w y)) y =
<i>(w.2 y).1.2@-i
-secEq (A B:U) (w:equiv A B) (x:A) : Id A (invEq A B w (w.1 x)) x =
+secEq (A B:U) (w:equiv A B) (x:A) : Path A (invEq A B w (w.1 x)) x =
<i> ((w.2 (w.1 x)).2 (x,<j>w.1 x)@i).1
-transEquiv (A B:U) (p:Id U A B) : equiv A B = (f,p)
+transEquiv (A B:U) (p:Path U A B) : equiv A B = (f,p)
where
f (x:A) : B = comp p x []
g (y:B) : A = comp (<i>p@-i) y []
- lem1 (x:A) : IdP p x (f x) = <i>comp (<j>p@(i/\j)) x [(i=0) -> <j>x]
- lem2 (y:B) : IdP p (g y) y = <i>comp (<j>p@(i\/-j)) y [(i=1) -> <j>y]
- lem3 (y:B) : Id B y (f (g y)) = <i>comp p (g y) [(i=0) -> lem2 y,(i=1) -> lem1 (g y)]
- lem4 (y:B) : IdP (<i>Id (p@i) (lem2 y@i) (lem1 (g y)@i)) (<j>g y) (lem3 y) =
+ lem1 (x:A) : PathP p x (f x) = <i>comp (<j>p@(i/\j)) x [(i=0) -> <j>x]
+ lem2 (y:B) : PathP p (g y) y = <i>comp (<j>p@(i\/-j)) y [(i=1) -> <j>y]
+ lem3 (y:B) : Path B y (f (g y)) = <i>comp p (g y) [(i=0) -> lem2 y,(i=1) -> lem1 (g y)]
+ lem4 (y:B) : PathP (<i>Path (p@i) (lem2 y@i) (lem1 (g y)@i)) (<j>g y) (lem3 y) =
<j i>fill p (g y) [(i=0) -> lem2 y,(i=1) -> lem1 (g y)] @ j
- lem5 (y:B) (x:A) (q:Id B y (f x)) : Id A (g y) x =
+ lem5 (y:B) (x:A) (q:Path B y (f x)) : Path A (g y) x =
<i>comp (<j>p@-j) (q@i) [(i=0) -> <j>lem2 y@-j,(i=1) -> <j>lem1 x@-j]
- lem6 (y:B) (x:A) (q:Id B y (f x)) : IdP (<i>Id (p@i) (lem2 y@i) (lem1 x@i)) (lem5 y x q) q =
+ lem6 (y:B) (x:A) (q:Path B y (f x)) : PathP (<i>Path (p@i) (lem2 y@i) (lem1 x@i)) (lem5 y x q) q =
<j i>fill (<j>p@-j) (q@i) [(i=0) -> <k>lem2 y@-k,(i=1) -> <k>lem1 x@-k] @ -j
- lem7 (y:B) (x:A) (q:Id B y (f x)) : IdP (<i>Id B y (f (lem5 y x q@i))) (lem3 y) q =
+ lem7 (y:B) (x:A) (q:Path B y (f x)) : PathP (<i>Path B y (f (lem5 y x q@i))) (lem3 y) q =
<j i>comp p (lem5 y x q@i/\j) [(i=0) -> lem2 y, (i=1) -> lem1 (lem5 y x q@j),(j=0) -> <k>lem4 y@k@i,(j=1) -> <k>lem6 y x q@k@i]
lem8 (y:B) : fiber A B f y = (g y,lem3 y)
- lem9 (y:B) (z:fiber A B f y) : Id (fiber A B f y) (lem8 y) z =
+ lem9 (y:B) (z:fiber A B f y) : Path (fiber A B f y) (lem8 y) z =
<i>(lem5 y z.1 z.2@i,lem7 y z.1 z.2@i)
p (y:B) : isContr (fiber A B f y) = (lem8 y,lem9 y)
-lemSinglContr (A:U) (a:A) : isContr ((x:A) * Id A a x) =
- ((a,refl A a),\ (z:(x:A) * Id A a x) -> contrSingl A a z.1 z.2)
+lemSinglContr (A:U) (a:A) : isContr ((x:A) * Path A a x) =
+ ((a,refl A a),\ (z:(x:A) * Path A a x) -> contrSingl A a z.1 z.2)
idEquiv (A:U) : equiv A A = (\ (x:A) -> x, lemSinglContr A)
-transEquiv (A X:U) (p:Id U A X) : equiv A X =
+transEquiv (A X:U) (p:Path U A X) : equiv A X =
substTrans U (equiv A) A X p (idEquiv A)
transDelta (A:U) : equiv A A = transEquiv A A (<i>A)
-transEquivToId (A B : U) (w : equiv A B) : Id U A B =
+transEquivToPath (A B : U) (w : equiv A B) : Path U A B =
<i> glue B [ (i = 1) -> (B,eB)
, (i = 0) -> (A,w) ]
where eB : equiv B B = transDelta B
-eqToEq (A B : U) (p : Id U A B)
- : Id (Id U A B) (transEquivToId A B (transEquiv A B p)) p
+eqToEq (A B : U) (p : Path U A B)
+ : Path (Path U A B) (transEquivToPath A B (transEquiv A B p)) p
= <j i> let e : equiv A B = transEquiv A B p
f : equiv B B = transDelta B
Ai : U = p@i
, (i = 1) -> (B,f)
, (j = 1) -> (p@i,g)]
-transIdFun (A B : U) (w : equiv A B)
- : Id (A -> B) w.1 (transEquiv A B (transEquivToId A B w)).1 =
+transPathFun (A B : U) (w : equiv A B)
+ : Path (A -> B) w.1 (transEquiv A B (transEquivToPath A B w)).1 =
<i> \ (a:A) ->
let b : B = w.1 a
u : A = comp (<j>A) a []
- q : Id B (w.1 u) b = <i>w.1 (comp (<j>A) a [(i=1) -> <j>a])
+ q : Path B (w.1 u) b = <i>w.1 (comp (<j>A) a [(i=1) -> <j>a])
in comp (<j> B)
(comp (<j> B) (comp (<j> B) (comp (<j> B) (w.1 u) [(i=0)->q]) [(i=0)-><k>b]) [(i=0)-><k>b]) [(i=0)-><k>b]
-idToId (A B : U) (w : equiv A B)
- : Id (equiv A B) (transEquiv A B (transEquivToId A B w)) w
- = equivLemma A B (transEquiv A B (transEquivToId A B w)) w
- (<i>transIdFun A B w@-i)
+idToPath (A B : U) (w : equiv A B)
+ : Path (equiv A B) (transEquiv A B (transEquivToPath A B w)) w
+ = equivLemma A B (transEquiv A B (transEquivToPath A B w)) w
+ (<i>transPathFun A B w@-i)
-- The grad lemma
lemIso (A B : U) (f : A -> B) (g : B -> A)
- (s : (y : B) -> Id B (f (g y)) y)
- (t : (x : A) -> Id A (g (f x)) x)
- (y : B) (x0 x1 : A) (p0 : Id B y (f x0)) (p1 : Id B y (f x1)) :
- Id (fiber A B f y) (x0,p0) (x1,p1) = <i> (p @ i,sq1 @ i)
+ (s : (y : B) -> Path B (f (g y)) y)
+ (t : (x : A) -> Path A (g (f x)) x)
+ (y : B) (x0 x1 : A) (p0 : Path B y (f x0)) (p1 : Path B y (f x1)) :
+ Path (fiber A B f y) (x0,p0) (x1,p1) = <i> (p @ i,sq1 @ i)
where
- rem0 : Id A (g y) x0 =
+ rem0 : Path A (g y) x0 =
<i> comp (<k> A) (g (p0 @ i)) [ (i = 1) -> t x0, (i = 0) -> <k> g y ]
- rem1 : Id A (g y) x1 =
+ rem1 : Path A (g y) x1 =
<i> comp (<k> A) (g (p1 @ i)) [ (i = 1) -> t x1, (i = 0) -> <k> g y ]
- p : Id A x0 x1 =
+ p : Path A x0 x1 =
<i> comp (<k> A) (g y) [ (i = 0) -> rem0
, (i = 1) -> rem1 ]
, (j = 0) -> s y ]
gradLemma (A B : U) (f : A -> B) (g : B -> A)
- (s : (y : B) -> Id B (f (g y)) y)
- (t : (x : A) -> Id A (g (f x)) x) : isEquiv A B f =
+ (s : (y : B) -> Path B (f (g y)) y)
+ (t : (x : A) -> Path A (g (f x)) x) : isEquiv A B f =
\(y:B) -> ((g y,<i>s y@-i),\ (z:fiber A B f y) ->
lemIso A B f g s t y (g y) z.1 (<i>s y@-i) z.2)
-isoId (A B : U) (f : A -> B) (g : B -> A)
- (s : (y : B) -> Id B (f (g y)) y)
- (t : (x : A) -> Id A (g (f x)) x) : Id U A B =
+isoPath (A B : U) (f : A -> B) (g : B -> A)
+ (s : (y : B) -> Path B (f (g y)) y)
+ (t : (x : A) -> Path A (g (f x)) x) : Path U A B =
<i> glue B [ (i = 0) -> (A,f,gradLemma A B f g s t)
, (i = 1) -> (B,idfun B,idIsEquiv B) ]
data gTrunc (A : U)
= inc (a : A)
- | squashC (a b : gTrunc A) (p q : Id (gTrunc A) a b)
- (r s: Id (Id (gTrunc A) a b) p q) <i j k>
+ | squashC (a b : gTrunc A) (p q : Path (gTrunc A) a b)
+ (r s: Path (Path (gTrunc A) a b) p q) <i j k>
[ (i=0) -> r @ j @ k
, (i=1) -> s @ j @ k
, (j=0) -> p @ k
, (k=0) -> a
, (k=1) -> b]
-gTr (A:U) (a b : gTrunc A) (p q : Id (gTrunc A) a b)
- (r s: Id (Id (gTrunc A) a b) p q) :
- Id (Id (Id (gTrunc A) a b) p q) r s =
+gTr (A:U) (a b : gTrunc A) (p q : Path (gTrunc A) a b)
+ (r s: Path (Path (gTrunc A) a b) p q) :
+ Path (Path (Path (gTrunc A) a b) p q) r s =
<i j k> squashC{gTrunc A} a b p q r s@ i @ j @k
gTruncRec (A B : U) (bG : groupoid B) (f : A -> B) : gTrunc A -> B = split
(<m n> gTruncRec A B bG f (s @ m @ n)) @ i @ j @ k
lem1 (A:U) (P:A -> U) (gP:(x:A) -> groupoid (P x)) (a :A) :
- (s:Id (Id A a a) (refl A a) (refl A a))
- (t:Id (Id (Id A a a) (refl A a) (refl A a)) (refl (Id A a a) (refl A a)) s)
- (a1 b1:P a) (p1 q1: Id (P a) a1 b1)
- (r1 : Id (Id (P a) a1 b1) p1 q1) (s1 : IdP (<i>IdP (<j>P (s@i@j)) a1 b1) p1 q1) ->
- IdP (<i>IdP (<j>IdP (<k>P (t@i@j@k)) a1 b1) p1 q1) r1 s1 =
- J (Id (Id A a a) (refl A a) (refl A a)) (refl (Id A a a) (refl A a))
- (\ (s:Id (Id A a a) (refl A a) (refl A a))
- (t:Id (Id (Id A a a) (refl A a) (refl A a)) (refl (Id A a a) (refl A a)) s) ->
- (a1 b1 :P a) (p1 q1: Id (P a) a1 b1)
- (r1 : Id (Id (P a) a1 b1) p1 q1) (s1 : IdP (<i>IdP (<j>P (s@i@j)) a1 b1) p1 q1) ->
- IdP (<i>IdP (<j>IdP (<k>P (t@i@j@k)) a1 b1) p1 q1) r1 s1) (gP a)
+ (s:Path (Path A a a) (refl A a) (refl A a))
+ (t:Path (Path (Path A a a) (refl A a) (refl A a)) (refl (Path A a a) (refl A a)) s)
+ (a1 b1:P a) (p1 q1: Path (P a) a1 b1)
+ (r1 : Path (Path (P a) a1 b1) p1 q1) (s1 : PathP (<i>PathP (<j>P (s@i@j)) a1 b1) p1 q1) ->
+ PathP (<i>PathP (<j>PathP (<k>P (t@i@j@k)) a1 b1) p1 q1) r1 s1 =
+ J (Path (Path A a a) (refl A a) (refl A a)) (refl (Path A a a) (refl A a))
+ (\ (s:Path (Path A a a) (refl A a) (refl A a))
+ (t:Path (Path (Path A a a) (refl A a) (refl A a)) (refl (Path A a a) (refl A a)) s) ->
+ (a1 b1 :P a) (p1 q1: Path (P a) a1 b1)
+ (r1 : Path (Path (P a) a1 b1) p1 q1) (s1 : PathP (<i>PathP (<j>P (s@i@j)) a1 b1) p1 q1) ->
+ PathP (<i>PathP (<j>PathP (<k>P (t@i@j@k)) a1 b1) p1 q1) r1 s1) (gP a)
lem (A:U) (P:A -> U) (gP:(x:A) -> groupoid (P x))
- (a :A) : (b:A) (p q:Id A a b) (r s:Id (Id A a b) p q) (t:Id (Id (Id A a b) p q) r s)
- (a1:P a) (b1:P b) (p1: IdP (<i>P (p@i)) a1 b1) (q1: IdP (<i>P (q@i)) a1 b1)
- (r1 : IdP (<i>IdP (<j>P (r@i@j)) a1 b1) p1 q1) (s1 : IdP (<i>IdP (<j>P (s@i@j)) a1 b1) p1 q1) ->
- IdP (<i>IdP (<j>IdP (<k>P (t@i@j@k)) a1 b1) p1 q1) r1 s1 =
- J A a (\ (b:A) (p :Id A a b) -> (q:Id A a b) (r s:Id (Id A a b) p q) (t:Id (Id (Id A a b) p q) r s)
- (a1:P a) (b1:P b) (p1: IdP (<i>P (p@i)) a1 b1) (q1: IdP (<i>P (q@i)) a1 b1)
- (r1 : IdP (<i>IdP (<j>P (r@i@j)) a1 b1) p1 q1) (s1 : IdP (<i>IdP (<j>P (s@i@j)) a1 b1) p1 q1) ->
- IdP (<i>IdP (<j>IdP (<k>P (t@i@j@k)) a1 b1) p1 q1) r1 s1) rem
+ (a :A) : (b:A) (p q:Path A a b) (r s:Path (Path A a b) p q) (t:Path (Path (Path A a b) p q) r s)
+ (a1:P a) (b1:P b) (p1: PathP (<i>P (p@i)) a1 b1) (q1: PathP (<i>P (q@i)) a1 b1)
+ (r1 : PathP (<i>PathP (<j>P (r@i@j)) a1 b1) p1 q1) (s1 : PathP (<i>PathP (<j>P (s@i@j)) a1 b1) p1 q1) ->
+ PathP (<i>PathP (<j>PathP (<k>P (t@i@j@k)) a1 b1) p1 q1) r1 s1 =
+ J A a (\ (b:A) (p :Path A a b) -> (q:Path A a b) (r s:Path (Path A a b) p q) (t:Path (Path (Path A a b) p q) r s)
+ (a1:P a) (b1:P b) (p1: PathP (<i>P (p@i)) a1 b1) (q1: PathP (<i>P (q@i)) a1 b1)
+ (r1 : PathP (<i>PathP (<j>P (r@i@j)) a1 b1) p1 q1) (s1 : PathP (<i>PathP (<j>P (s@i@j)) a1 b1) p1 q1) ->
+ PathP (<i>PathP (<j>PathP (<k>P (t@i@j@k)) a1 b1) p1 q1) r1 s1) rem
where
- rem : (q:Id A a a) (r s:Id (Id A a a) (refl A a) q) (t:Id (Id (Id A a a) (refl A a) q) r s)
- (a1:P a) (b1:P a) (p1: Id (P a) a1 b1) (q1: IdP (<i>P (q@i)) a1 b1)
- (r1 : IdP (<i>IdP (<j>P (r@i@j)) a1 b1) p1 q1) (s1 : IdP (<i>IdP (<j>P (s@i@j)) a1 b1) p1 q1) ->
- IdP (<i>IdP (<j>IdP (<k>P (t@i@j@k)) a1 b1) p1 q1) r1 s1 =
- J (Id A a a) (refl A a) (\ (q:Id A a a) (r:Id (Id A a a) (refl A a) q) ->
- (s:Id (Id A a a) (refl A a) q) (t:Id (Id (Id A a a) (refl A a) q) r s)
- (a1:P a) (b1:P a) (p1: Id (P a) a1 b1) (q1: IdP (<i>P (q@i)) a1 b1)
- (r1 : IdP (<i>IdP (<j>P (r@i@j)) a1 b1) p1 q1) (s1 : IdP (<i>IdP (<j>P (s@i@j)) a1 b1) p1 q1) ->
- IdP (<i>IdP (<j>IdP (<k>P (t@i@j@k)) a1 b1) p1 q1) r1 s1) (lem1 A P gP a)
+ rem : (q:Path A a a) (r s:Path (Path A a a) (refl A a) q) (t:Path (Path (Path A a a) (refl A a) q) r s)
+ (a1:P a) (b1:P a) (p1: Path (P a) a1 b1) (q1: PathP (<i>P (q@i)) a1 b1)
+ (r1 : PathP (<i>PathP (<j>P (r@i@j)) a1 b1) p1 q1) (s1 : PathP (<i>PathP (<j>P (s@i@j)) a1 b1) p1 q1) ->
+ PathP (<i>PathP (<j>PathP (<k>P (t@i@j@k)) a1 b1) p1 q1) r1 s1 =
+ J (Path A a a) (refl A a) (\ (q:Path A a a) (r:Path (Path A a a) (refl A a) q) ->
+ (s:Path (Path A a a) (refl A a) q) (t:Path (Path (Path A a a) (refl A a) q) r s)
+ (a1:P a) (b1:P a) (p1: Path (P a) a1 b1) (q1: PathP (<i>P (q@i)) a1 b1)
+ (r1 : PathP (<i>PathP (<j>P (r@i@j)) a1 b1) p1 q1) (s1 : PathP (<i>PathP (<j>P (s@i@j)) a1 b1) p1 q1) ->
+ PathP (<i>PathP (<j>PathP (<k>P (t@i@j@k)) a1 b1) p1 q1) r1 s1) (lem1 A P gP a)
T:U = (A:U) (P:A -> U) (gP:(x:A) -> groupoid (P x))
- (a b:A) (p q:Id A a b) (r s:Id (Id A a b) p q) (t:Id (Id (Id A a b) p q) r s)
- (a1:P a) (b1:P b) (p1: IdP (<i>P (p@i)) a1 b1) (q1: IdP (<i>P (q@i)) a1 b1)
- (r1 : IdP (<i>IdP (<j>P (r@i@j)) a1 b1) p1 q1) (s1 : IdP (<i>IdP (<j>P (s@i@j)) a1 b1) p1 q1) ->
- IdP (<i>IdP (<j>IdP (<k>P (t@i@j@k)) a1 b1) p1 q1) r1 s1
+ (a b:A) (p q:Path A a b) (r s:Path (Path A a b) p q) (t:Path (Path (Path A a b) p q) r s)
+ (a1:P a) (b1:P b) (p1: PathP (<i>P (p@i)) a1 b1) (q1: PathP (<i>P (q@i)) a1 b1)
+ (r1 : PathP (<i>PathP (<j>P (r@i@j)) a1 b1) p1 q1) (s1 : PathP (<i>PathP (<j>P (s@i@j)) a1 b1) p1 q1) ->
+ PathP (<i>PathP (<j>PathP (<k>P (t@i@j@k)) a1 b1) p1 q1) r1 s1
gTruncElim1 (lem:T) (A : U)
(B : (gTrunc A) -> U)
(bG : (z:gTrunc A) -> groupoid (B z))
(f : (x:A) -> B (inc x)) (z:gTrunc A) -> B z = gTruncElim1 lem
-univG (A B:U) (bG:groupoid B) : Id U ((gTrunc A) -> B) (A -> B) =
- isoId (gTrunc A -> B) (A -> B) F G s t
+univG (A B:U) (bG:groupoid B) : Path U ((gTrunc A) -> B) (A -> B) =
+ isoPath (gTrunc A -> B) (A -> B) F G s t
where
F (h : gTrunc A -> B) (a: A) : B = h (inc a)
G : (A -> B) -> gTrunc A -> B = gTruncRec A B bG
- s (h : A -> B) : Id (A -> B) (F (G h)) h = <i> \ (x:A) -> h x
+ s (h : A -> B) : Path (A -> B) (F (G h)) h = <i> \ (x:A) -> h x
- t (h : gTrunc A -> B) : Id (gTrunc A -> B) (G (F h)) h = <i> \ (z:gTrunc A) -> rem z @ i
+ t (h : gTrunc A -> B) : Path (gTrunc A -> B) (G (F h)) h = <i> \ (z:gTrunc A) -> rem z @ i
where
- P (z:gTrunc A) : U = Id B (G (F h) z) (h z)
+ P (z:gTrunc A) : U = Path B (G (F h) z) (h z)
tP (z : gTrunc A) : groupoid (P z) = setGroupoid (P z) (bG (G (F h) z) (h z))
\r
import prelude\r
\r
-hedbergLemma (A: U) (a b:A) (f : (x : A) -> Id A a x -> Id A a x) (p : Id A a b) :\r
+hedbergLemma (A: U) (a b:A) (f : (x : A) -> Path A a x -> Path A a x) (p : Path A a b) :\r
Square A a a a b (refl A a) p (f a (refl A a)) (f b p) = \r
comp (<i> Square A a a a (p @ i) (<_> a) (<j> p @ i /\ j)\r
(f a (<_> a)) (f (p @ i) (<j> p @ i /\ j)))\r
(<i> f a (<_> a)) []\r
\r
-hedbergStable (A : U) (a b : A) (h : (x : A) -> stable (Id A a x))\r
- (p q : Id A a b) : Id (Id A a b) p q =\r
+hedbergStable (A : U) (a b : A) (h : (x : A) -> stable (Path A a x))\r
+ (p q : Path A a b) : Path (Path A a b) p q =\r
<j i> comp (<_> A) a [ (j = 0) -> rem2 @ i\r
, (j = 1) -> rem3 @ i\r
, (i = 0) -> r\r
, (i = 1) -> rem4 @ j]\r
where \r
- ra : Id A a a = <_> a \r
- rem1 (x : A) : exConst (Id A a x) = stableConst (Id A a x) (h x)\r
- f (x : A) : Id A a x -> Id A a x = (rem1 x).1\r
- fIsConst (x : A) : const (Id A a x) (f x) = (rem1 x).2\r
+ ra : Path A a a = <_> a \r
+ rem1 (x : A) : exConst (Path A a x) = stableConst (Path A a x) (h x)\r
+ f (x : A) : Path A a x -> Path A a x = (rem1 x).1\r
+ fIsConst (x : A) : const (Path A a x) (f x) = (rem1 x).2\r
rem4 : Square A a a b b ra (refl A b) (f b p) (f b q) = fIsConst b p q\r
- r : Id A a a = f a ra\r
+ r : Path A a a = f a ra\r
rem2 : Square A a a a b ra p r (f b p) = hedbergLemma A a b f p\r
rem3 : Square A a a a b ra q r (f b q) = hedbergLemma A a b f q\r
\r
-hedbergS (A:U) (h : (a x:A) -> stable (Id A a x)) : set A =\r
+hedbergS (A:U) (h : (a x:A) -> stable (Path A a x)) : set A =\r
\(a b : A) -> hedbergStable A a b (h a)\r
\r
hedberg (A : U) (h : discrete A) : set A =\r
- \(a b : A) -> hedbergStable A a b (\(b : A) -> decStable (Id A a b) (h a b))\r
+ \(a b : A) -> hedbergStable A a b (\(b : A) -> decStable (Path A a b) (h a b))\r
\r
\r
\r
\r
-- Alternative version:\r
\r
--- hedbergLemma (A: U) (f : (a b : A) -> Id A a b -> Id A a b) (a :A) :\r
--- (b : A) (p : Id A a b) ->\r
--- Id (Id A a b) (compId A a a b (f a a (refl A a)) p) (f a b p) =\r
--- J A a (\ (b:A) (p:Id A a b) -> Id (Id A a b) (compId A a a b (f a a (refl A a)) p) (f a b p))\r
--- (refl (Id A a a) (f a a (refl A a)))\r
+-- hedbergLemma (A: U) (f : (a b : A) -> Path A a b -> Path A a b) (a :A) :\r
+-- (b : A) (p : Path A a b) ->\r
+-- Path (Path A a b) (compPath A a a b (f a a (refl A a)) p) (f a b p) =\r
+-- J A a (\ (b:A) (p:Path A a b) -> Path (Path A a b) (compPath A a a b (f a a (refl A a)) p) (f a b p))\r
+-- (refl (Path A a a) (f a a (refl A a)))\r
\r
--- hedberg (A : U) (h : discrete A) : set A = \(a b : A) (p q : Id A a b) ->\r
--- let rem1 (x y : A) : exConst (Id A x y) = decConst (Id A x y) (h x y)\r
+-- hedberg (A : U) (h : discrete A) : set A = \(a b : A) (p q : Path A a b) ->\r
+-- let rem1 (x y : A) : exConst (Path A x y) = decConst (Path A x y) (h x y)\r
\r
--- f (x y : A) : Id A x y -> Id A x y = (rem1 x y).1\r
+-- f (x y : A) : Path A x y -> Path A x y = (rem1 x y).1\r
\r
--- fIsConst (x y : A) : const (Id A x y) (f x y) = (rem1 x y).2\r
+-- fIsConst (x y : A) : const (Path A x y) (f x y) = (rem1 x y).2\r
\r
--- r : Id A a a = f a a (refl A a)\r
+-- r : Path A a a = f a a (refl A a)\r
\r
--- rem2 : Id (Id A a b) (compId A a a b r p) (f a b p) = hedbergLemma A f a b p\r
+-- rem2 : Path (Path A a b) (compPath A a a b r p) (f a b p) = hedbergLemma A f a b p\r
\r
--- rem3 : Id (Id A a b) (compId A a a b r q) (f a b q) = hedbergLemma A f a b q\r
+-- rem3 : Path (Path A a b) (compPath A a a b r q) (f a b q) = hedbergLemma A f a b q\r
\r
--- rem4 : Id (Id A a b) (f a b p) (f a b q) = fIsConst a b p q\r
+-- rem4 : Path (Path A a b) (f a b p) (f a b q) = fIsConst a b p q\r
\r
--- rem5 : Id (Id A a b) (compId A a a b r p) (compId A a a b r q) =\r
--- compDown (Id A a b) (compId A a a b r p) (f a b p) (compId A a a b r q)\r
+-- rem5 : Path (Path A a b) (compPath A a a b r p) (compPath A a a b r q) =\r
+-- compDown (Path A a b) (compPath A a a b r p) (f a b p) (compPath A a a b r q)\r
-- (f a b q) rem2 rem3 rem4\r
-- in lemSimpl A a a b r p q rem5\r
backTurn (l: loopS1) : loopS1 = compS1 l invLoop
-compInv (A:U) (a:A) : (x:A) (p:Id A a x) -> Id (Id A x x) (<_>x) (compId A x a x (<i>p@-i) p) =
- J A a (\ (x:A) (p:Id A a x) -> Id (Id A x x) (<_>x) (compId A x a x (<i>p@-i) p)) rem
- where rem : Id (Id A a a) (<_>a) (<i>comp (<_>A) a [(i=0) -> <_>a,(i=1) -> <_>a]) =
+compInv (A:U) (a:A) : (x:A) (p:Path A a x) -> Path (Path A x x) (<_>x) (compPath A x a x (<i>p@-i) p) =
+ J A a (\ (x:A) (p:Path A a x) -> Path (Path A x x) (<_>x) (compPath A x a x (<i>p@-i) p)) rem
+ where rem : Path (Path A a a) (<_>a) (<i>comp (<_>A) a [(i=0) -> <_>a,(i=1) -> <_>a]) =
<j i>comp (<_>A) a [(j=0) -> <_>a,(i=0) -> <_>a,(i=1) -> <_>a]
-compInvS1 : Id loopS1 (refl S1 base) (compS1 invLoop loop1) = compInv S1 base base loop1
+compInvS1 : Path loopS1 (refl S1 base) (compS1 invLoop loop1) = compInv S1 base base loop1
-compInv (A:U) (a b:A) (q:Id A a b) : (x:A) (p:Id A b x) -> Id (Id A a b) q (compId A a x b (compId A a b x q p) (<i>p@-i)) =
- J A b (\ (x:A) (p:Id A b x) -> Id (Id A a b) q (compId A a x b (compId A a b x q p) (<i>p@-i))) rem
- where rem : Id (Id A a b) q
+compInv (A:U) (a b:A) (q:Path A a b) : (x:A) (p:Path A b x) -> Path (Path A a b) q (compPath A a x b (compPath A a b x q p) (<i>p@-i)) =
+ J A b (\ (x:A) (p:Path A b x) -> Path (Path A a b) q (compPath A a x b (compPath A a b x q p) (<i>p@-i))) rem
+ where rem : Path (Path A a b) q
(<i>comp (<_>A) (comp (<_>A) (q@i) [(i=0) -> <_>a,(i=1) -> <_>b]) [(i=0) -> <_>a,(i=1) -> <_>b]) =
<j i>comp (<_>A) (comp (<_>A) (q@i) [(j=0) -> <_>q@i,(i=0) -> <_>a,(i=1) -> <_>b]) [(j=0) -> <_>q@i,(i=0) -> <_>a,(i=1) -> <_>b]
-- it is equal to the identity function
-lemTransC (A:U) (a:A) : Id A (transC A a) a = <i>comp (<_>A) a [(i=1) -> <_>a]
+lemTransC (A:U) (a:A) : Path A (transC A a) a = <i>comp (<_>A) a [(i=1) -> <_>a]
lemFib1 (A:U) (F G : A -> U) (a:A) (fa : F a -> G a) :
- (x:A) (p : Id A a x) -> (fx : F x -> G x) ->
- Id U (Id (F a -> G x) (\ (u:F a) -> subst A G a x p (fa u)) (\ (u:F a) -> fx (subst A F a x p u)))
- (IdP (<i>F (p@i) -> G (p@i)) fa fx) =
- J A a (\ (x:A) (p : Id A a x) -> (fx : F x -> G x) ->
- Id U (Id (F a -> G x) (\ (u:F a) -> subst A G a x p (fa u)) (\ (u:F a) -> fx (subst A F a x p u)))
- (IdP (<i>F (p@i) -> G (p@i)) fa fx)) rem
+ (x:A) (p : Path A a x) -> (fx : F x -> G x) ->
+ Path U (Path (F a -> G x) (\ (u:F a) -> subst A G a x p (fa u)) (\ (u:F a) -> fx (subst A F a x p u)))
+ (PathP (<i>F (p@i) -> G (p@i)) fa fx) =
+ J A a (\ (x:A) (p : Path A a x) -> (fx : F x -> G x) ->
+ Path U (Path (F a -> G x) (\ (u:F a) -> subst A G a x p (fa u)) (\ (u:F a) -> fx (subst A F a x p u)))
+ (PathP (<i>F (p@i) -> G (p@i)) fa fx)) rem
where
- rem (ga : F a -> G a) : Id U (Id (F a -> G a) (\ (u:F a) -> transC (G a) (fa u)) (\ (u:F a) -> ga (transC (F a) u)))
- (Id (F a -> G a) fa ga) =
- <j>Id (F a -> G a) (\ (u:F a) -> lemTransC (G a) (fa u)@j) (\ (u : F a) -> ga (lemTransC (F a) u@j))
+ rem (ga : F a -> G a) : Path U (Path (F a -> G a) (\ (u:F a) -> transC (G a) (fa u)) (\ (u:F a) -> ga (transC (F a) u)))
+ (Path (F a -> G a) fa ga) =
+ <j>Path (F a -> G a) (\ (u:F a) -> lemTransC (G a) (fa u)@j) (\ (u : F a) -> ga (lemTransC (F a) u@j))
-- special case
-corFib1 (A:U) (F G : A -> U) (a:A) (fa ga : F a -> G a) (p:Id A a a)
- (h : (u:F a) -> Id (G a) (subst A G a a p (fa u)) (ga (subst A F a a p u))) : IdP (<i>F (p@i) -> G (p@i)) fa ga =
+corFib1 (A:U) (F G : A -> U) (a:A) (fa ga : F a -> G a) (p:Path A a a)
+ (h : (u:F a) -> Path (G a) (subst A G a a p (fa u)) (ga (subst A F a a p u))) : PathP (<i>F (p@i) -> G (p@i)) fa ga =
comp (lemFib1 A F G a fa a p ga) (<i>\ (u:F a) -> h u@i) []
-compIdL (A:U) (a b:A) (p : Id A a b) : Id (Id A a b) p (compId A a b b p (<_>b)) =
+compPathL (A:U) (a b:A) (p : Path A a b) : Path (Path A a b) p (compPath A a b b p (<_>b)) =
<j i>comp (<_>A) (p @ i) [(i=0) -> <_> a, (i = 1) -> <_>b, (j=0) -> <_>(p@i) ]
-lemFib2 (A:U) (F:A->U) (a b:A) (p:Id A a b) (u:F a) :
- (c:A) (q:Id A b c) -> Id (F c) (subst A F b c q (subst A F a b p u)) (subst A F a c (compId A a b c p q) u) =
- J A b (\ (c:A) (q:Id A b c) -> Id (F c) (subst A F b c q (subst A F a b p u)) (subst A F a c (compId A a b c p q) u))
+lemFib2 (A:U) (F:A->U) (a b:A) (p:Path A a b) (u:F a) :
+ (c:A) (q:Path A b c) -> Path (F c) (subst A F b c q (subst A F a b p u)) (subst A F a c (compPath A a b c p q) u) =
+ J A b (\ (c:A) (q:Path A b c) -> Path (F c) (subst A F b c q (subst A F a b p u)) (subst A F a c (compPath A a b c p q) u))
rem
where
- rem1 : Id (F b) (subst A F a b p u) (subst A F b b (<_>b) (subst A F a b p u)) = <i>lemTransC (F b) (subst A F a b p u)@-i
- rem2 : Id (F b) (subst A F a b p u) (subst A F a b (compId A a b b p (<_>b)) u) =
- <i>subst A F a b (compIdL A a b p@i) u
- rem : Id (F b) (subst A F b b (<_>b) (subst A F a b p u)) (subst A F a b (compId A a b b p (<_>b)) u) =
- comp (<i>Id (F b) (rem1@i) (rem2@i)) (<_>subst A F a b p u) []
-
-testIsoId (A B : U) (f : A -> B) (g : B -> A)
- (s : (y : B) -> Id B (f (g y)) y)
- (t : (x : A) -> Id A (g (f x)) x) (a:A) : Id B (f a) (trans A B (isoId A B f g s t) a) =
+ rem1 : Path (F b) (subst A F a b p u) (subst A F b b (<_>b) (subst A F a b p u)) = <i>lemTransC (F b) (subst A F a b p u)@-i
+ rem2 : Path (F b) (subst A F a b p u) (subst A F a b (compPath A a b b p (<_>b)) u) =
+ <i>subst A F a b (compPathL A a b p@i) u
+ rem : Path (F b) (subst A F b b (<_>b) (subst A F a b p u)) (subst A F a b (compPath A a b b p (<_>b)) u) =
+ comp (<i>Path (F b) (rem1@i) (rem2@i)) (<_>subst A F a b p u) []
+
+testIsoPath (A B : U) (f : A -> B) (g : B -> A)
+ (s : (y : B) -> Path B (f (g y)) y)
+ (t : (x : A) -> Path A (g (f x)) x) (a:A) : Path B (f a) (trans A B (isoPath A B f g s t) a) =
<i>comp (<_>B) (comp (<_>B) (f a) [(i=0) -> <_> f a]) [(i=0) -> <_> f a]
testH (n:Z) : Z = subst S1 helix base base loop1 n
-testHelix : Id (Z->Z) sucZ (subst S1 helix base base loop1) =
- <i>\(n:Z) -> testIsoId Z Z sucZ predZ sucpredZ predsucZ n@i
+testHelix : Path (Z->Z) sucZ (subst S1 helix base base loop1) =
+ <i>\(n:Z) -> testIsoPath Z Z sucZ predZ sucpredZ predsucZ n@i
-encode (x:S1) (p:Id S1 base x) : helix x = subst S1 helix base x p zeroZ
+encode (x:S1) (p:Path S1 base x) : helix x = subst S1 helix base x p zeroZ
itLoop : nat -> loopS1 = split
zero -> triv
inl n -> itLoopNeg n
inr n -> itLoop n
-lemItNeg : (n:nat) -> Id loopS1 (transport (<i>Id S1 base (loop{S1} @ i)) (loopIt (inl n))) (loopIt (sucZ (inl n))) = split
+lemItNeg : (n:nat) -> Path loopS1 (transport (<i>Path S1 base (loop{S1} @ i)) (loopIt (inl n))) (loopIt (sucZ (inl n))) = split
zero -> lemInv S1 base base loop1
suc n -> lemCompInv S1 base base base (itLoopNeg n) invLoop
l0 : loopS1 = <i>base
l1 : loopS1 = oneTurn l0
-test1ItPos (n:nat) : U = Id loopS1 (loopIt (sucZ (inr n))) (oneTurn ((loopIt (inr n))))
+test1ItPos (n:nat) : U = Path loopS1 (loopIt (sucZ (inr n))) (oneTurn ((loopIt (inr n))))
-lem1ItPos : (n:nat) -> Id loopS1 (loopIt (sucZ (inr n))) (oneTurn ((loopIt (inr n)))) = split
+lem1ItPos : (n:nat) -> Path loopS1 (loopIt (sucZ (inr n))) (oneTurn ((loopIt (inr n)))) = split
zero -> refl loopS1 l1
suc p -> <i>oneTurn (lem1ItPos p@i)
-test1ItNeg (n:nat) : U = Id loopS1 (loopIt (sucZ (inl n))) (oneTurn ((loopIt (inl n))))
+test1ItNeg (n:nat) : U = Path loopS1 (loopIt (sucZ (inl n))) (oneTurn ((loopIt (inl n))))
-lem1ItNeg : (n:nat) -> Id loopS1 (loopIt (sucZ (inl n))) (oneTurn (loopIt (inl n))) = split
+lem1ItNeg : (n:nat) -> Path loopS1 (loopIt (sucZ (inl n))) (oneTurn (loopIt (inl n))) = split
zero -> compInvS1
suc p -> compInv S1 base base (loopIt (inl p)) base invLoop
-lem1It : (n:Z) -> Id loopS1 (loopIt (sucZ n)) (oneTurn (loopIt n)) = split
+lem1It : (n:Z) -> Path loopS1 (loopIt (sucZ n)) (oneTurn (loopIt n)) = split
inl n -> lem1ItNeg n
inr n -> lem1ItPos n
-decode : (x:S1) -> helix x -> Id S1 base x = split
+decode : (x:S1) -> helix x -> Path S1 base x = split
base -> loopIt
loop @ i -> rem @ i
where T : U = Z -> loopS1
- G (x:S1) : U = Id S1 base x
- p : Id U T T = <j> helix (loop1@j) -> Id S1 base (loop1@j)
- rem2 (n:Z) : Id loopS1 (oneTurn (loopIt n)) (loopIt (sucZ n)) = <i>lem1It n@-i
- rem1 (n:Z) : Id loopS1 (subst S1 G base base loop1 (loopIt n)) (loopIt (subst S1 helix base base loop1 n)) =
- comp (<i> Id loopS1 (oneTurn (loopIt n)) (loopIt (testHelix@i n))) (rem2 n) []
- rem : IdP p loopIt loopIt = corFib1 S1 helix G base loopIt loopIt loop1 rem1
-
-encodeDecode (x:S1) (p : Id S1 base x) : Id (Id S1 base x) (decode x (encode x p)) p =
- transport (<i>Id (Id S1 base (p@i)) (decode (p@i) (encode (p@i) (<j>p@(i/\j)))) (<j>p@(i/\j))) (refl loopS1 triv)
-
--- encodeDecode : (c : S1) (p : Id S1 base c) -> Id (Id S1 base c) (decode c (encode c p)) p =
--- J S1 base (\ (c : S1) (p : Id S1 base c) -> Id (Id S1 base c) (decode c (encode c p)) p)
+ G (x:S1) : U = Path S1 base x
+ p : Path U T T = <j> helix (loop1@j) -> Path S1 base (loop1@j)
+ rem2 (n:Z) : Path loopS1 (oneTurn (loopIt n)) (loopIt (sucZ n)) = <i>lem1It n@-i
+ rem1 (n:Z) : Path loopS1 (subst S1 G base base loop1 (loopIt n)) (loopIt (subst S1 helix base base loop1 n)) =
+ comp (<i> Path loopS1 (oneTurn (loopIt n)) (loopIt (testHelix@i n))) (rem2 n) []
+ rem : PathP p loopIt loopIt = corFib1 S1 helix G base loopIt loopIt loop1 rem1
+
+encodeDecode (x:S1) (p : Path S1 base x) : Path (Path S1 base x) (decode x (encode x p)) p =
+ transport (<i>Path (Path S1 base (p@i)) (decode (p@i) (encode (p@i) (<j>p@(i/\j)))) (<j>p@(i/\j))) (refl loopS1 triv)
+
+-- encodeDecode : (c : S1) (p : Path S1 base c) -> Path (Path S1 base c) (decode c (encode c p)) p =
+-- J S1 base (\ (c : S1) (p : Path S1 base c) -> Path (Path S1 base c) (decode c (encode c p)) p)
-- (<_> (<_> base))
-lemTransOneTurn (n:nat) : Id Z (transport (<i>helix (loop1@i)) (inr n)) (inr (suc n)) =
+lemTransOneTurn (n:nat) : Path Z (transport (<i>helix (loop1@i)) (inr n)) (inr (suc n)) =
<i>inr (suc (comp (<_>nat) (comp (<_>nat) n [(i=1) -> <_>n]) [(i=1) -> <_>n]))
-lemTransBackTurn (n:nat) : Id Z (transport (<i>helix (loop1@-i)) (inl n)) (inl (suc n)) =
+lemTransBackTurn (n:nat) : Path Z (transport (<i>helix (loop1@-i)) (inl n)) (inl (suc n)) =
<i>inl (suc (comp (<_>nat) (comp (<_>nat) n [(i=1) -> <_>n]) [(i=1) -> <_>n]))
-corFib2 (u:Z) (l:loopS1) : Id Z (transport (<i>helix (oneTurn l@i)) u)
+corFib2 (u:Z) (l:loopS1) : Path Z (transport (<i>helix (oneTurn l@i)) u)
(transport (<i>helix (loop1@i)) (transport (<i>helix (l@i)) u)) =
<i>lemFib2 S1 helix base base l u base loop1@-i
-corFib3 (u:Z) (l:loopS1) : Id Z (transport (<i>helix (backTurn l@i)) u)
+corFib3 (u:Z) (l:loopS1) : Path Z (transport (<i>helix (backTurn l@i)) u)
(transport (<i>helix (loop1@-i)) (transport (<i>helix (l@i)) u)) =
<i>lemFib2 S1 helix base base l u base (<j>loop1@-j)@-i
-decodeEncodeBasePos : (n : nat) -> Id Z (transport (<x> helix (itLoop n @ x)) (inr zero)) (inr n) = split
+decodeEncodeBasePos : (n : nat) -> Path Z (transport (<x> helix (itLoop n @ x)) (inr zero)) (inr n) = split
zero -> <_> inr zero
- suc n -> comp (<j>Id Z (transport (<i>helix (oneTurn l@i)) (inr zero)) (lemTransOneTurn n@j)) rem3 []
+ suc n -> comp (<j>Path Z (transport (<i>helix (oneTurn l@i)) (inr zero)) (lemTransOneTurn n@j)) rem3 []
where l : loopS1 = itLoop n
- rem1 : Id Z (transport (<i> helix (l@i)) (inr zero)) (inr n) = decodeEncodeBasePos n
- rem2 : Id Z (transport (<i>helix (oneTurn l@i)) (inr zero))
+ rem1 : Path Z (transport (<i> helix (l@i)) (inr zero)) (inr n) = decodeEncodeBasePos n
+ rem2 : Path Z (transport (<i>helix (oneTurn l@i)) (inr zero))
(transport (<i>helix (loop1@i)) (transport (<i>helix (l@i)) (inr zero))) =
corFib2 (inr zero) l
- rem3 : Id Z (transport (<i>helix (oneTurn l@i)) (inr zero))
+ rem3 : Path Z (transport (<i>helix (oneTurn l@i)) (inr zero))
(transport (<i>helix (loop1@i)) (inr n)) =
- comp (<j>Id Z (transport (<i>helix (oneTurn l@i)) (inr zero))
+ comp (<j>Path Z (transport (<i>helix (oneTurn l@i)) (inr zero))
(transport (<i>helix (loop1@i)) (rem1@j))) rem2 []
-decodeEncodeBaseNeg : (n : nat) -> Id Z (transport (<x> helix (itLoopNeg n @ x)) (inr zero)) (inl n) = split
+decodeEncodeBaseNeg : (n : nat) -> Path Z (transport (<x> helix (itLoopNeg n @ x)) (inr zero)) (inl n) = split
zero -> <_> inl zero
- suc n -> comp (<j>Id Z (transport (<i>helix (backTurn l@i)) (inr zero)) (lemTransBackTurn n@j)) rem3 []
+ suc n -> comp (<j>Path Z (transport (<i>helix (backTurn l@i)) (inr zero)) (lemTransBackTurn n@j)) rem3 []
where l : loopS1 = itLoopNeg n
- rem1 : Id Z (transport (<i> helix (l@i)) (inr zero)) (inl n) = decodeEncodeBaseNeg n
- rem2 : Id Z (transport (<i>helix (backTurn l@i)) (inr zero))
+ rem1 : Path Z (transport (<i> helix (l@i)) (inr zero)) (inl n) = decodeEncodeBaseNeg n
+ rem2 : Path Z (transport (<i>helix (backTurn l@i)) (inr zero))
(transport (<i>helix (loop1@-i)) (transport (<i>helix (l@i)) (inr zero))) =
corFib3 (inr zero) l
- rem3 : Id Z (transport (<i>helix (backTurn l@i)) (inr zero))
+ rem3 : Path Z (transport (<i>helix (backTurn l@i)) (inr zero))
(transport (<i>helix (loop1@-i)) (inl n)) =
- comp (<j>Id Z (transport (<i>helix (backTurn l@i)) (inr zero))
+ comp (<j>Path Z (transport (<i>helix (backTurn l@i)) (inr zero))
(transport (<i>helix (loop1@-i)) (rem1@j))) rem2 []
-decodeEncodeBase : (n : Z) -> Id Z (encode base (decode base n)) n = split
+decodeEncodeBase : (n : Z) -> Path Z (encode base (decode base n)) n = split
inl n -> decodeEncodeBaseNeg n
inr n -> decodeEncodeBasePos n
-- the loop space of the circle is equal to Z
-loopS1equalsZ : Id U loopS1 Z =
- isoId loopS1 Z (encode base) (decode base) decodeEncodeBase (encodeDecode base)
+loopS1equalsZ : Path U loopS1 Z =
+ isoPath loopS1 Z (encode base) (decode base) decodeEncodeBase (encodeDecode base)
setLoop : set loopS1 = substInv U set loopS1 Z loopS1equalsZ ZSet
-- S1 is a groupoid
isGroupoidS1 : groupoid S1 = lem
where
- lem2 : (y : S1) -> set (Id S1 base y)
- = lemPropFib (\ (y:S1) -> set (Id S1 base y)) (\ (y:S1) -> setIsProp (Id S1 base y)) setLoop
+ lem2 : (y : S1) -> set (Path S1 base y)
+ = lemPropFib (\ (y:S1) -> set (Path S1 base y)) (\ (y:S1) -> setIsProp (Path S1 base y)) setLoop
- lem : (x y : S1) -> set (Id S1 x y)
- = lemPropFib (\ (x:S1) -> (y : S1) -> set (Id S1 x y)) pP lem2
+ lem : (x y : S1) -> set (Path S1 x y)
+ = lemPropFib (\ (x:S1) -> (y : S1) -> set (Path S1 x y)) pP lem2
where
- pP (x:S1) : prop ((y:S1) -> set (Id S1 x y)) =
- propPi S1 (\ (y:S1) -> set (Id S1 x y)) (\ (y:S1) -> setIsProp (Id S1 x y))
+ pP (x:S1) : prop ((y:S1) -> set (Path S1 x y)) =
+ propPi S1 (\ (y:S1) -> set (Path S1 x y)) (\ (y:S1) -> setIsProp (Path S1 x y))
-substInv (A : U) (P : A -> U) (a x : A) (p : Id A a x) : P x -> P a =
+substInv (A : U) (P : A -> U) (a x : A) (p : Path A a x) : P x -> P a =
subst A P x a (<i>p @ -i)
-funDepTr (A0 A1:U) (p:Id U A0 A1) (u0:A0) (u1:A1) :
- Id U (IdP p u0 u1) (Id A1 (transport p u0) u1) =
- <i> IdP (<l>p @ (i\/l)) (comp (<l>p @ (i/\l)) u0 [(i=0) -> <_>u0]) u1
+funDepTr (A0 A1:U) (p:Path U A0 A1) (u0:A0) (u1:A1) :
+ Path U (PathP p u0 u1) (Path A1 (transport p u0) u1) =
+ <i> PathP (<l>p @ (i\/l)) (comp (<l>p @ (i/\l)) u0 [(i=0) -> <_>u0]) u1
lemSetTorus (E : S1 -> S1 -> U) (sE : set (E base base))
(f : (y:S1) -> E base y) (g : (x:S1) -> E x base)
- (efg : Id (E base base) (f base) (g base)) : (x y:S1) -> E x y = split
+ (efg : Path (E base base) (f base) (g base)) : (x y:S1) -> E x y = split
base -> f
loop @ i -> lem2 @ i
where
G (y x:S1) : U = E x y
- lem1 : (y:S1) -> IdS S1 (G y) base base loop1 (f y) (f y) = lemPropFib P pP bP
+ lem1 : (y:S1) -> PathS S1 (G y) base base loop1 (f y) (f y) = lemPropFib P pP bP
where
- P (y:S1) : U = IdS S1 (G y) base base loop1 (f y) (f y)
+ P (y:S1) : U = PathS S1 (G y) base base loop1 (f y) (f y)
sbE : (y : S1) -> set (E base y)
= lemPropFib (\ (y:S1) -> set (E base y)) (\ (y:S1) -> setIsProp (E base y)) sE
pP (y:S1) : prop (P y) = rem3
where
- rem1 : Id U (P y) (Id (E base y) (subst S1 (G y) base base loop1 (f y)) (f y))
+ rem1 : Path U (P y) (Path (E base y) (subst S1 (G y) base base loop1 (f y)) (f y))
= funDepTr (G y base) (G y base) (<j>G y (loop{S1} @ j)) (f y) (f y)
- rem2 : prop (Id (E base y) (subst S1 (G y) base base loop1 (f y)) (f y))
+ rem2 : prop (Path (E base y) (subst S1 (G y) base base loop1 (f y)) (f y))
= sbE y (subst S1 (G y) base base loop1 (f y)) (f y)
rem3 : prop (P y)
- = substInv U prop (P y) (Id (E base y) (subst S1 (G y) base base loop1 (f y)) (f y)) rem1 rem2
+ = substInv U prop (P y) (Path (E base y) (subst S1 (G y) base base loop1 (f y)) (f y)) rem1 rem2
- lem2 : IdS S1 (G base) base base loop1 (g base) (g base)
+ lem2 : PathS S1 (G base) base base loop1 (g base) (g base)
= <j> g (loop1 @ j)
bP : P base
- = substInv (E base base) (\ (u:E base base) -> IdS S1 (G base) base base loop1 u u) (f base) (g base) efg lem2
+ = substInv (E base base) (\ (u:E base base) -> PathS S1 (G base) base base loop1 u u) (f base) (g base) efg lem2
- lem2 : IdS S1 F base base loop1 f f = <j> \ (y:S1) -> (lem1 y) @ j
+ lem2 : PathS S1 F base base loop1 f f = <j> \ (y:S1) -> (lem1 y) @ j
-- commutativity of mult, at last
-idL : (x : S1) -> Id S1 (mult base x) x = split
+idL : (x : S1) -> Path S1 (mult base x) x = split
base -> refl S1 base
loop @ i -> <j> loop1 @ i
-multCom : (x y : S1) -> Id S1 (mult x y) (mult y x) =
+multCom : (x y : S1) -> Path S1 (mult x y) (mult y x) =
lemSetTorus E sE idL g efg
where
- E (x y: S1) : U = Id S1 (mult x y) (mult y x)
+ E (x y: S1) : U = Path S1 (mult x y) (mult y x)
sE : set (E base base) = isGroupoidS1 base base
g (x : S1) : E x base = inv S1 (mult base x) (mult x base) (idL x)
- efg : Id (E base base) (idL base) (g base) = refl (E base base) (idL base)
+ efg : Path (E base base) (idL base) (g base) = refl (E base base) (idL base)
-- associativity
-multAssoc (x :S1) : (y z : S1) -> Id S1 (mult x (mult y z)) (mult (mult x y) z) =
+multAssoc (x :S1) : (y z : S1) -> Path S1 (mult x (mult y z)) (mult (mult x y) z) =
lemSetTorus E sE f g efg
where
- E (y z : S1) : U = Id S1 (mult x (mult y z)) (mult (mult x y) z)
+ E (y z : S1) : U = Path S1 (mult x (mult y z)) (mult (mult x y) z)
sE : set (E base base) = isGroupoidS1 x x
f (z : S1) : E base z = rem
where
- rem1 : Id S1 (mult base z) z = multCom base z
+ rem1 : Path S1 (mult base z) z = multCom base z
- rem : Id S1 (mult x (mult base z)) (mult x z) = <i> mult x (rem1 @ i)
+ rem : Path S1 (mult x (mult base z)) (mult x z) = <i> mult x (rem1 @ i)
g (y : S1) : E y base = refl S1 (mult x y)
- efg : Id (E base base) (f base) (g base) = refl (E base base) (f base)
+ efg : Path (E base base) (f base) (g base) = refl (E base base) (f base)
-- inverse law
(lemPropFib (P base) (pP base) bP)
invLaw : (x y : S1) ->
- Id (Id S1 (mult x y) (mult x y)) (refl S1 (mult x y))
- (compId S1 (mult x y) (mult y x) (mult x y) (multCom x y) (multCom y x)) = lemPropRel P pP bP
+ Path (Path S1 (mult x y) (mult x y)) (refl S1 (mult x y))
+ (compPath S1 (mult x y) (mult y x) (mult x y) (multCom x y) (multCom y x)) = lemPropRel P pP bP
where
P (x y : S1) : U
- = Id (Id S1 (mult x y) (mult x y)) (refl S1 (mult x y))
- (compId S1 (mult x y) (mult y x) (mult x y) (multCom x y) (multCom y x))
+ = Path (Path S1 (mult x y) (mult x y)) (refl S1 (mult x y))
+ (compPath S1 (mult x y) (mult y x) (mult x y) (multCom x y) (multCom y x))
pP (x y : S1) : prop (P x y) =
isGroupoidS1 (mult x y) (mult x y) (refl S1 (mult x y))
- (compId S1 (mult x y) (mult y x) (mult x y) (multCom x y) (multCom y x))
+ (compPath S1 (mult x y) (mult y x) (mult x y) (multCom x y) (multCom y x))
bP : P base base =
- comp (<i>Id (Id S1 base base) (refl S1 base) (<j>comp (<_>S1) base [(i=0) -> refl S1 base, (j=0) -> refl S1 base, (j=1) -> refl S1 base]))
- (refl (Id S1 base base) (refl S1 base)) []
+ comp (<i>Path (Path S1 base base) (refl S1 base) (<j>comp (<_>S1) base [(i=0) -> refl S1 base, (j=0) -> refl S1 base, (j=1) -> refl S1 base]))
+ (refl (Path S1 base base) (refl S1 base)) []
-- the multiplication is invertible
multIsEquiv : (x:S1) -> isEquiv S1 S1 (mult x) = lemPropFib P pP bP
where P (x:S1) : U = isEquiv S1 S1 (mult x)
pP (x:S1) : prop (P x) = propIsEquiv S1 S1 (mult x)
- rem : Id (S1 -> S1) (idfun S1) (mult base) = <i>\ (x:S1) -> idL x @ -i
+ rem : Path (S1 -> S1) (idfun S1) (mult base) = <i>\ (x:S1) -> idL x @ -i
bP : P base = subst (S1->S1) (isEquiv S1 S1) (idfun S1) (mult base) rem (idIsEquiv S1)
-- inverse of multiplication by x
invS1 (x:S1) : S1 = invMult x base
-lemInvS1 : Id S1 (invS1 base) base = <i> comp (<_>S1) (comp (<_>S1) base [(i=1) -> refl S1 base]) [(i=1) -> refl S1 base]
+lemInvS1 : Path S1 (invS1 base) base = <i> comp (<_>S1) (comp (<_>S1) base [(i=1) -> refl S1 base]) [(i=1) -> refl S1 base]
-loopInvS1 : U = Id S1 (invS1 base) (invS1 base)
+loopInvS1 : U = Path S1 (invS1 base) (invS1 base)
-rePar (l: loopInvS1) : loopS1 = transport (<i>Id S1 (lemInvS1@i) (lemInvS1@i)) l
+rePar (l: loopInvS1) : loopS1 = transport (<i>Path S1 (lemInvS1@i) (lemInvS1@i)) l
test2 : Z = winding (rePar (<i>invS1 (loop2@i)))
-- EVAL: inl (suc zero) Time: 1m26.400s
-- This reduces to "zero"
test1 : nat = comp (<_> nat) zero []
-test2 : Id hnat nzero (comp (<_> hnat) nzero []) =
+test2 : Path hnat nzero (comp (<_> hnat) nzero []) =
fill (<_> hnat) nzero []
toNat : hnat -> nat = split
zero -> nzero
suc n -> nsuc (fromNat n)
-toNatK : (n : hnat) -> Id hnat (fromNat (toNat n)) n = split
+toNatK : (n : hnat) -> Path hnat (fromNat (toNat n)) n = split
nzero -> <_> nzero
nsuc n -> <i> nsuc (toNatK n @ i)
-fromNatK : (n : nat) -> Id nat (toNat (fromNat n)) n = split
+fromNatK : (n : nat) -> Path nat (toNat (fromNat n)) n = split
zero -> <_> zero
suc n -> <i> suc (fromNatK n @ i)
-hnatEqNat : Id U hnat nat =
- isoId hnat nat toNat fromNat fromNatK toNatK
+hnatEqNat : Path U hnat nat =
+ isoPath hnat nat toNat fromNat fromNatK toNatK
-- This is zero
test3 : nat = trans hnat nat hnatEqNat test0
test4 : hnat = trans nat hnat (<i> hnatEqNat @ -i) zero
-- This is "hComp (hnat) (hComp (hnat) nzero []) []"
-test5 : hnat = trans hnat hnat (compId U hnat nat hnat hnatEqNat (<i> hnatEqNat @ -i)) nzero
+test5 : hnat = trans hnat hnat (compPath U hnat nat hnat hnatEqNat (<i> hnatEqNat @ -i)) nzero
hnatSet : set hnat = subst U set nat hnat (<i> hnatEqNat @ -i) natSet
-- shorthand for nat x nat
nat2 : U = and nat nat
-natlemma (a b c d : nat) : Id nat (add (add a b) (add c d)) (add (add a d) (add c b)) =
- let rem : Id nat (add a (add b (add c d))) (add a (add d (add c b))) =
+natlemma (a b c d : nat) : Path nat (add (add a b) (add c d)) (add (add a d) (add c b)) =
+ let rem : Path nat (add a (add b (add c d))) (add a (add d (add c b))) =
<i> add a (add_comm3 b c d @ i)
in <i> comp (<_> nat) (rem @ i) [ (i = 0) -> assocAdd a b (add c d)
, (i = 1) -> assocAdd a d (add c b) ]
rel : eqrel nat2 = (r,rem)
where
r : hrel nat2 = \(x y : nat2) ->
- (Id nat (add x.1 y.2) (add x.2 y.1),natSet (add x.1 y.2) (add x.2 y.1))
+ (Path nat (add x.1 y.2) (add x.2 y.1),natSet (add x.1 y.2) (add x.2 y.1))
rem : iseqrel nat2 r = ((rem1,rem2),rem3)
where
rem1 : istrans nat2 r =
\(x y z : nat2)
- (h1 : Id nat (add x.1 y.2) (add x.2 y.1))
- (h2 : Id nat (add y.1 z.2) (add y.2 z.1)) ->
- let rem : Id nat (add (add x.1 y.2) (add y.1 z.2)) (add (add x.2 y.1) (add y.2 z.1)) =
+ (h1 : Path nat (add x.1 y.2) (add x.2 y.1))
+ (h2 : Path nat (add y.1 z.2) (add y.2 z.1)) ->
+ let rem : Path nat (add (add x.1 y.2) (add y.1 z.2)) (add (add x.2 y.1) (add y.2 z.1)) =
<i> add (h1 @ i) (h2 @ i)
- rem1 : Id nat (add (add x.1 y.2) (add y.1 z.2)) (add (add x.1 z.2) (add y.1 y.2)) =
+ rem1 : Path nat (add (add x.1 y.2) (add y.1 z.2)) (add (add x.1 z.2) (add y.1 y.2)) =
natlemma x.1 y.2 y.1 z.2
- rem2 : Id nat (add (add x.2 y.1) (add y.2 z.1)) (add (add x.2 z.1) (add y.2 y.1)) =
+ rem2 : Path nat (add (add x.2 y.1) (add y.2 z.1)) (add (add x.2 z.1) (add y.2 y.1)) =
natlemma x.2 y.1 y.2 z.1
- rem3 : Id nat (add (add x.2 z.1) (add y.2 y.1)) (add (add x.2 z.1) (add y.1 y.2)) =
+ rem3 : Path nat (add (add x.2 z.1) (add y.2 y.1)) (add (add x.2 z.1) (add y.1 y.2)) =
<i> add (add x.2 z.1) (add_comm y.2 y.1 @ i)
- rem4 : Id nat (add (add x.2 y.1) (add y.2 z.1)) (add (add x.2 z.1) (add y.1 y.2)) =
+ rem4 : Path nat (add (add x.2 y.1) (add y.2 z.1)) (add (add x.2 z.1) (add y.1 y.2)) =
<i> comp (<_> nat) (add (add x.2 z.1) (add y.2 y.1)) [ (i = 0) -> <j> rem2 @ -j
, (i = 1) -> rem3 ]
- rem5 : Id nat (add (add x.1 z.2) (add y.1 y.2)) (add (add x.2 z.1) (add y.1 y.2)) =
+ rem5 : Path nat (add (add x.1 z.2) (add y.1 y.2)) (add (add x.2 z.1) (add y.1 y.2)) =
<i> comp (<_> nat) (rem @ i) [ (i = 0) -> rem1, (i = 1) -> rem4 ]
in natcancelr (add x.1 z.2) (add x.2 z.1) (add y.1 y.2) rem5
rem2 : isrefl nat2 r = \(x : nat2) -> add_comm x.1 x.2
- rem3 : issymm nat2 r = \(x y : nat2) (h : Id nat (add x.1 y.2) (add x.2 y.1)) ->
- let rem : Id nat (add x.2 y.1) (add y.2 x.1) =
+ rem3 : issymm nat2 r = \(x y : nat2) (h : Path nat (add x.1 y.2) (add x.2 y.1)) ->
+ let rem : Path nat (add x.2 y.1) (add y.2 x.1) =
<i> comp (<_> nat) (add x.1 y.2) [ (i = 0) -> h
, (i = 1) -> add_comm x.1 y.2 ]
in <i> comp (<_> nat) (add x.2 y.1) [ (i = 0) -> add_comm x.2 y.1
p <i> []
propNoPoints : prop NoPoints = split
- p @ i -> let rem : (b : NoPoints) -> Id NoPoints (p{NoPoints} @ i) b = split
+ p @ i -> let rem : (b : NoPoints) -> Path NoPoints (p{NoPoints} @ i) b = split
p @ j -> <k> p{NoPoints} @ (i /\ -k) \/ (j /\ k)
in rem
point0 : NoPoints = p{NoPoints} @ 0
point1 : NoPoints = p{NoPoints} @ 1
-p' : Id NoPoints point0 point1 = <i> p{NoPoints} @ i
+p' : Path NoPoints point0 point1 = <i> p{NoPoints} @ i
f1 : NoPoints -> Unit = split
p @ i -> tt
f2 : Unit -> NoPoints = split
tt -> point0
-test : Id U NoPoints Unit =
- isoId NoPoints Unit f1 f2 rem1 rem2
+test : Path U NoPoints Unit =
+ isoPath NoPoints Unit f1 f2 rem1 rem2
where
- rem1 : (y : Unit) -> Id Unit (f1 (f2 y)) y = split
+ rem1 : (y : Unit) -> Path Unit (f1 (f2 y)) y = split
tt -> <i> tt
- rem2 : (x : NoPoints) -> Id NoPoints (f2 (f1 x)) x = split
+ rem2 : (x : NoPoints) -> Path NoPoints (f2 (f1 x)) x = split
p @ i -> <j> p{NoPoints} @ j /\ i
-fext (A B : U) (f g : A -> B) (h : (x : A) -> Id B (f x) (g x)) :
- Id (A -> B) f g = <j> (\(x : A) -> htpy x (p{NoPoints} @ j))
+fext (A B : U) (f g : A -> B) (h : (x : A) -> Path B (f x) (g x)) :
+ Path (A -> B) f g = <j> (\(x : A) -> htpy x (p{NoPoints} @ j))
where htpy (x : A) : NoPoints -> B = split
p @ i -> h x @ i
-- First definition of injectivity, informally: if two elements f a0, f a1 are
-- equal in B, then a0, a1 must be equal in A.
inj0 (A B : U) (f : A -> B) (sA : set A) (sB : set B) : U
- = (a0 a1 : A) -> Id B (f a0) (f a1) -> Id A a0 a1
+ = (a0 a1 : A) -> Path B (f a0) (f a1) -> Path A a0 a1
-- Second definition of injectivity, informally: for any b in B, there are
-- only one elment a in A such that f a is equal to b.
inj1 (A B : U) (f : A -> B) (sA : set A) (sB : set B) : U
- = (b : B) -> prop ((a : A) * Id B (f a) b)
+ = (b : B) -> prop ((a : A) * Path B (f a) b)
-- A map from the first to the second definition.
inj01 (A B : U) (f : A -> B) (sA : set A) (sB : set B) : inj0 A B f sA sB ->
inj1 A B f sA sB
- = \ (i0 : inj0 A B f sA sB) (b : B) (c d : (a : A) * Id B (f a) b) -> let
+ = \ (i0 : inj0 A B f sA sB) (b : B) (c d : (a : A) * Path B (f a) b) -> let
F (a : A) : U
- = Id B (f a) b
+ = Path B (f a) b
pF (a : A) : prop (F a)
= sB (f a) b
- p : Id B (f c.1) (f d.1)
+ p : Path B (f c.1) (f d.1)
= <i> comp (<j> B) (c.2 @ i) [ (i = 0) -> <j> f c.1
, (i = 1) -> <j> d.2 @ -j ]
- q : Id A c.1 d.1
+ q : Path A c.1 d.1
= i0 c.1 d.1 p
in
lemSig A F pF c d q
-- A map from the second to the first definition.
inj10 (A B : U) (f : A -> B) (sA : set A) (sB : set B) : inj1 A B f sA sB ->
inj0 A B f sA sB
- = \ (i1 : inj1 A B f sA sB) (a0 a1 : A) (p : Id B (f a0) (f a1)) -> let
- c : (a : A) * Id B (f a) (f a1)
+ = \ (i1 : inj1 A B f sA sB) (a0 a1 : A) (p : Path B (f a0) (f a1)) -> let
+ c : (a : A) * Path B (f a) (f a1)
= (a0, p)
- d : (a : A) * Id B (f a) (f a1)
+ d : (a : A) * Path B (f a) (f a1)
= (a1, <i> f a1)
- q : Id ((a : A) * Id B (f a) (f a1)) c d
+ q : Path ((a : A) * Path B (f a) (f a1)) c d
= i1 (f a1) c d
- fst : ((a : A) * Id B (f a) (f a1)) -> A
- = \ (x : (a : A) * Id B (f a) (f a1)) -> x.1
+ fst : ((a : A) * Path B (f a) (f a1)) -> A
+ = \ (x : (a : A) * Path B (f a) (f a1)) -> x.1
in
<i> fst (q @ i)
prop_inj0 (A B : U) (f : A -> B) (sA : set A) (sB : set B)
: prop (inj0 A B f sA sB)
= let
- c (a0 a1 : A) : prop (Id B (f a0) (f a1) -> Id A a0 a1)
+ c (a0 a1 : A) : prop (Path B (f a0) (f a1) -> Path A a0 a1)
= let
- P : Id B (f a0) (f a1) -> U
- = \ (_ : Id B (f a0) (f a1)) -> Id A a0 a1
- h : (x : Id B (f a0) (f a1)) -> prop (P x)
- = \ (_ : Id B (f a0) (f a1)) -> sA a0 a1
+ P : Path B (f a0) (f a1) -> U
+ = \ (_ : Path B (f a0) (f a1)) -> Path A a0 a1
+ h : (x : Path B (f a0) (f a1)) -> prop (P x)
+ = \ (_ : Path B (f a0) (f a1)) -> sA a0 a1
in
- propPi (Id B (f a0) (f a1)) P h
- d (a0 : A) : prop ((a1 : A) -> Id B (f a0) (f a1) -> Id A a0 a1)
+ propPi (Path B (f a0) (f a1)) P h
+ d (a0 : A) : prop ((a1 : A) -> Path B (f a0) (f a1) -> Path A a0 a1)
= let
P : A -> U
- = \ (a1 : A) -> ( Id B (f a0) (f a1) -> Id A a0 a1 )
+ = \ (a1 : A) -> ( Path B (f a0) (f a1) -> Path A a0 a1 )
h : (a1 : A) -> prop (P a1)
= \ (a1 : A) -> c a0 a1
in
e : prop (inj0 A B f sA sB)
= let
P : A -> U
- = \ (a0 : A) -> ( (a1 : A) -> Id B (f a0) (f a1) -> Id A a0 a1 )
- h : (a0 : A) -> prop ( (a1 : A) -> Id B (f a0) (f a1) -> Id A a0 a1 )
+ = \ (a0 : A) -> ( (a1 : A) -> Path B (f a0) (f a1) -> Path A a0 a1 )
+ h : (a0 : A) -> prop ( (a1 : A) -> Path B (f a0) (f a1) -> Path A a0 a1 )
= \ (a0 : A) -> d a0
in
propPi A P h
prop (inj1 A B f sA sB)
= let
P : B -> U
- = \ (b : B) -> (a : A) * Id B (f a) b
+ = \ (b : B) -> (a : A) * Path B (f a) b
Q : B -> U
= \ (b : B) -> prop (P b)
h : (b : B) -> prop (Q b)
-- A proof that two propositions with maps between them can be identified with
-- each other
-propId (A B : U) (f : A -> B) (g : B -> A) (pA : prop A) (pB : prop B) :
- Id U A B
- = isoId A B f g (\ (b : B) -> pB (f (g b)) b) (\ (a : A) -> pA (g (f a)) a)
+propPath (A B : U) (f : A -> B) (g : B -> A) (pA : prop A) (pB : prop B) :
+ Path U A B
+ = isoPath A B f g (\ (b : B) -> pB (f (g b)) b) (\ (a : A) -> pA (g (f a)) a)
-- A proof that the two definitions of injectivity can be identified with each
-- other
-injId (A B : U) (f : A -> B) (sA : set A) (sB : set B) :
- Id U (inj0 A B f sA sB) (inj1 A B f sA sB)
- = propId (inj0 A B f sA sB) (inj1 A B f sA sB)
+injPath (A B : U) (f : A -> B) (sA : set A) (sB : set B) :
+ Path U (inj0 A B f sA sB) (inj1 A B f sA sB)
+ = propPath (inj0 A B f sA sB) (inj1 A B f sA sB)
(inj01 A B f sA sB) (inj10 A B f sA sB)
(prop_inj0 A B f sA sB) (prop_inj1 A B f sA sB)
\ No newline at end of file
zero -> inl zero\r
suc n -> inr n\r
\r
-sucpredZ : (x : Z) -> Id Z (sucZ (predZ x)) x = split\r
+sucpredZ : (x : Z) -> Path Z (sucZ (predZ x)) x = split\r
inl u -> refl Z (inl u)\r
inr v -> lem v\r
where\r
- lem : (u : nat) -> Id Z (sucZ (predZ (inr u))) (inr u) = split\r
+ lem : (u : nat) -> Path Z (sucZ (predZ (inr u))) (inr u) = split\r
zero -> refl Z (inr zero)\r
suc n -> refl Z (inr (suc n))\r
\r
-predsucZ : (x : Z) -> Id Z (predZ (sucZ x)) x = split\r
+predsucZ : (x : Z) -> Path Z (predZ (sucZ x)) x = split\r
inl u -> lem u\r
where\r
- lem : (u : nat) -> Id Z (predZ (sucZ (inl u))) (inl u) = split\r
+ lem : (u : nat) -> Path Z (predZ (sucZ (inl u))) (inl u) = split\r
zero -> refl Z (inl zero)\r
suc n -> refl Z (inl (suc n))\r
inr v -> refl Z (inr v)\r
\r
\r
-sucIdZ : Id U Z Z = isoId Z Z sucZ predZ sucpredZ predsucZ\r
+sucPathZ : Path U Z Z = isoPath Z Z sucZ predZ sucpredZ predsucZ\r
\r
-- We can transport along the proof forward and backwards:\r
-testOneZ : Z = transport sucIdZ zeroZ\r
-testNOneZ : Z = transport (<i> sucIdZ @ - i) zeroZ\r
+testOneZ : Z = transport sucPathZ zeroZ\r
+testNOneZ : Z = transport (<i> sucPathZ @ - i) zeroZ\r
\r
ZSet : set Z = hedberg Z (orDisc nat nat natDec natDec)\r
| zeroP <i> [(i=0) -> pos zero, (i=1) -> neg zero]
-- Nice version of the zero constructor:
-zeroZ : Id int (pos zero) (neg zero) = <i> zeroP {int} @ i
+zeroZ : Path int (pos zero) (neg zero) = <i> zeroP {int} @ i
sucInt : int -> int = split
pos n -> pos (suc n)
inl n -> neg (suc n)
inr n -> pos n
-toZK : (a : Z) -> Id Z (toZ (fromZ a)) a = split
+toZK : (a : Z) -> Path Z (toZ (fromZ a)) a = split
inl n -> refl Z (inl n)
inr n -> refl Z (inr n)
-fromZK : (a : int) -> Id int (fromZ (toZ a)) a = split
+fromZK : (a : int) -> Path int (fromZ (toZ a)) a = split
pos n -> refl int (pos n)
neg n -> rem n
- where rem : (n : nat) -> Id int (fromZ (toZ (neg n))) (neg n) = split
+ where rem : (n : nat) -> Path int (fromZ (toZ (neg n))) (neg n) = split
zero -> zeroZ
suc m -> refl int (neg (suc m))
zeroP @ i -> <j> zeroZ @ i /\ j
-isoIntZ : Id U Z int = isoId Z int fromZ toZ fromZK toZK
+isoIntZ : Path U Z int = isoPath Z int fromZ toZ fromZK toZK
intSet : set int = subst U set Z int isoIntZ ZSet
-- a concrete instance
-T : U = Id int (pos zero) (pos zero)
+T : U = Path int (pos zero) (pos zero)
p0 : T = refl int (pos zero)
-p1 : T = compId int (pos zero) (neg zero) (pos zero) zeroZ (<i>zeroZ@-i)
+p1 : T = compPath int (pos zero) (neg zero) (pos zero) zeroZ (<i>zeroZ@-i)
-test0 : Id (Id Z (inr zero) (inr zero)) (refl Z (inr zero)) (refl Z (inr zero)) =
+test0 : Path (Path Z (inr zero) (inr zero)) (refl Z (inr zero)) (refl Z (inr zero)) =
ZSet (inr zero) (inr zero) (refl Z (inr zero)) (refl Z (inr zero))
-- Tests for normal forms:
-test1 : Id T p0 p1 = intSet (pos zero) (pos zero) p0 p1
-test2 : Id T p0 p0 = intSet (pos zero) (pos zero) p0 p0
+test1 : Path T p0 p1 = intSet (pos zero) (pos zero) p0 p1
+test2 : Path T p0 p0 = intSet (pos zero) (pos zero) p0 p0
-ntest1 : Id T p0 p1 = <i1 i2> comp (<_> int) (pos zero) [ (i1 = 0) -> <i3> pos zero, (i1 = 1) -> <i3> comp (<_> int) (zeroP {int} @ (i2 /\ i3)) [ (i2 = 0) -> <i4> pos zero, (i2 = 1) -> <i4> zeroP {int} @ (-i4 /\ i3), (i3 = 0) -> <i4> pos zero, (i3 = 1) -> <i4> comp (<_> int) (zeroP {int} @ i2) [ (i2 = 0) -> <i5> pos zero, (i2 = 1) -> <i5> zeroP {int} @ (-i4 \/ -i5), (i4 = 0) -> <i5> zeroP {int} @ i2 ] ], (i2 = 0) -> <i3> pos zero, (i2 = 1) -> <i3> pos zero ]
+ntest1 : Path T p0 p1 = <i1 i2> comp (<_> int) (pos zero) [ (i1 = 0) -> <i3> pos zero, (i1 = 1) -> <i3> comp (<_> int) (zeroP {int} @ (i2 /\ i3)) [ (i2 = 0) -> <i4> pos zero, (i2 = 1) -> <i4> zeroP {int} @ (-i4 /\ i3), (i3 = 0) -> <i4> pos zero, (i3 = 1) -> <i4> comp (<_> int) (zeroP {int} @ i2) [ (i2 = 0) -> <i5> pos zero, (i2 = 1) -> <i5> zeroP {int} @ (-i4 \/ -i5), (i4 = 0) -> <i5> zeroP {int} @ i2 ] ], (i2 = 0) -> <i3> pos zero, (i2 = 1) -> <i3> pos zero ]
-ntest2 : Id T p0 p0 = <i1 i2> comp (<_> int) (pos zero) [ (i1 = 0) -> <i3> pos zero, (i1 = 1) -> <i3> pos zero, (i2 = 0) -> <i3> pos zero, (i2 = 1) -> <i3> pos zero ]
+ntest2 : Path T p0 p0 = <i1 i2> comp (<_> int) (pos zero) [ (i1 = 0) -> <i3> pos zero, (i1 = 1) -> <i3> pos zero, (i2 = 0) -> <i3> pos zero, (i2 = 1) -> <i3> pos zero ]
data I = zero | one | seg <i> [(i = 0) -> zero, (i = 1) -> one]
-- Proof of funext from the interval
-fext (A B : U) (f g : A -> B) (p : (x : A) -> Id B (f x) (g x)) :
- Id (A -> B) f g = <j> (\(x : A) -> htpy x (seg{I} @ j))
+fext (A B : U) (f g : A -> B) (p : (x : A) -> Path B (f x) (g x)) :
+ Path (A -> B) f g = <j> (\(x : A) -> htpy x (seg{I} @ j))
where htpy (x : A) : I -> B = split
zero -> f x
one -> g x
fromUnit : Unit -> I = split
tt -> zero
-toUnitK : (a : Unit) -> Id Unit (toUnit (fromUnit a)) a = split
+toUnitK : (a : Unit) -> Path Unit (toUnit (fromUnit a)) a = split
tt -> <i> tt
-fromUnitK : (a : I) -> Id I (fromUnit (toUnit a)) a = split
+fromUnitK : (a : I) -> Path I (fromUnit (toUnit a)) a = split
zero -> <i> zero
one -> <i> seg {I} @ i
seg @ i -> <j> seg {I} @ i /\ j
-unitEqI : Id U Unit I = isoId Unit I fromUnit toUnit fromUnitK toUnitK
+unitEqI : Path U Unit I = isoPath Unit I fromUnit toUnit fromUnitK toUnitK
propI : prop I = subst U prop Unit I unitEqI propUnit
setI : set I = subst U set Unit I unitEqI setUnit
-T : U = Id I zero zero
+T : U = Path I zero zero
p0 : T = refl I zero
test : T = propI zero zero
cons x xs -> cons (f x) (map A B f xs)
lem (A B C:U) (f:A->B) (g:B -> C) :
- (xs:list A) -> Id (list C) (map B C g (map A B f xs)) (map A C (\ (x:A) -> g (f x)) xs) = split
+ (xs:list A) -> Path (list C) (map B C g (map A B f xs)) (map A C (\ (x:A) -> g (f x)) xs) = split
nil -> <i>nil
cons x xs -> <i>cons (g (f x)) (lem A B C f g xs@i)
-funId (A:U) (x:A) : A = x
+funPath (A:U) (x:A) : A = x
-lem1 (A:U) : (xs:list A) -> Id (list A) (map A A (funId A) xs) xs = split
+lem1 (A:U) : (xs:list A) -> Path (list A) (map A A (funPath A) xs) xs = split
nil -> <i>nil
cons x xs -> <i>cons x (lem1 A xs@i)
nil -> nil
cons x xs -> append A (reverse A xs) (cons x nil)
-lem2 (A:U) : (xs:list A) -> Id (list A) (append A xs nil) xs = split
+lem2 (A:U) : (xs:list A) -> Path (list A) (append A xs nil) xs = split
nil -> <i>nil
cons x xs -> <i>cons x (lem2 A xs@i)
-assoc (A:U) : (xs ys zs : list A) -> Id (list A) (append A (append A xs ys) zs) (append A xs (append A ys zs)) = split
+assoc (A:U) : (xs ys zs : list A) -> Path (list A) (append A (append A xs ys) zs) (append A xs (append A ys zs)) = split
nil -> \ (ys zs:list A) -> <i>append A ys zs
cons x xs -> \ (ys zs:list A) -> <i>cons x (assoc A xs ys zs@i)
{-
-lem4 (A:U) : (xs ys:list A) -> Id (list A) (reverse A (append A xs ys)) (append A (reverse A ys) (reverse A xs)) = split
+lem4 (A:U) : (xs ys:list A) -> Path (list A) (reverse A (append A xs ys)) (append A (reverse A ys) (reverse A xs)) = split
nil -> \ (ys:list A) -> <i>lem2 A (reverse A ys)@-i
cons x xs -> \ (ys:list A) -> <i>comp (list A) (append A (lem4 A xs ys@i) (cons x nil))
[(i=1) -> assoc A (reverse A ys) (reverse A xs) (cons x nil)]
-lem5 (A:U) : (xs:list A) -> Id (list A) (reverse A (reverse A xs)) xs = split
+lem5 (A:U) : (xs:list A) -> Path (list A) (reverse A (reverse A xs)) xs = split
nil -> <i>nil
cons x xs -> <i>comp (list A) (lem4 A (reverse A xs) (cons x nil)@i) [(i=1) -> <j>cons x (lem5 A xs@j)]
base -> bP
loop @ i -> (lemPropF S1 P pP base base loop1 bP bP) @ i
-idL : (x : S1) -> Id S1 (mult base x) x = split
+idL : (x : S1) -> Path S1 (mult base x) x = split
base -> refl S1 base
loop @ i -> <j> loop1 @ i
multIsEquiv : (x:S1) -> isEquiv S1 S1 (mult x) = lemPropFib P pP bP
where P (x:S1) : U = isEquiv S1 S1 (mult x)
pP (x:S1) : prop (P x) = propIsEquiv S1 S1 (mult x)
- rem : Id (S1 -> S1) (idfun S1) (mult base) = <i>\ (x:S1) -> idL x @ -i
+ rem : Path (S1 -> S1) (idfun S1) (mult base) = <i>\ (x:S1) -> idL x @ -i
bP : P base = subst (S1->S1) (isEquiv S1 S1) (idfun S1) (mult base) rem (idIsEquiv S1)
-- inverse of multiplication by x
pt0 : S1 = mapOnPath S1 S1 invS1 base base loop2@0
test1 : S1 = mapOnPath S1 S1 invS1 base base loop2@1
-test2 : Id S1 pt0 pt0 = mapOnPath S1 S1 invS1 base base loop2
+test2 : Path S1 pt0 pt0 = mapOnPath S1 S1 invS1 base base loop2
-ntest2 : Id S1 pt0 pt0 =
+ntest2 : Path S1 pt0 pt0 =
<i> comp (<_>S1) (comp (<_>S1) (comp (<_>S1) (comp (<_>S1) (comp (<_>S1) (comp (<_>S1) (comp (<_>S1) (comp (<_>S1) (comp (<_>S1) (comp (<_>S1) base []) []) []) [ (i = 0) -> <j> comp (<_>S1) (comp (<_>S1) (comp (<_>S1) base []) []) [ (j = 1) -> <k> comp (<_>S1) (comp (<_>S1) base []) [] ], (i = 1) -> <j> comp (<_>S1) (comp (<_>S1) (comp (<_>S1) (comp (<_>S1) base [ (j = 0) -> <k> comp (<_>S1) base [ (k = 0) -> <l> base ], (j = 1) -> <k> comp (<_>S1) base [ (k = 0) -> <l> base ] ]) [ (j = 0) -> <k> comp (<_>S1) (comp (<_>S1) base []) [ (k = 0) -> <l> comp (<_>S1) base [] ], (j = 1) -> <k> comp (<_>S1) (comp (<_>S1) base []) [ (k = 0) -> <l> comp (<_>S1) base [] ] ]) [ (j = 0) -> <k> comp (<_>S1) (comp (<_>S1) (comp (<_>S1) base []) []) [ (k = 0) -> <l> comp (<_>S1) (comp (<_>S1) base []) [] ], (j = 1) -> <k> comp (<_>S1) (comp (<_>S1) (comp (<_>S1) base []) []) [ (k = 0) -> <l> comp (<_>S1) (comp (<_>S1) base []) [] ] ]) [ (j = 0) -> <k> comp (<_>S1) (comp (<_>S1) (comp (<_>S1) base []) []) [], (j = 1) -> <k> comp (<_>S1) (comp (<_>S1) (comp (<_>S1) (comp (<_>S1) base []) []) []) [ (k = 0) -> <l> comp (<_>S1) (comp (<_>S1) (comp (<_>S1) base []) []) [] ] ] ]) [ (i = 0) -> <j> comp (<_>S1) (comp (<_>S1) base []) [ (j = 1) -> <k> comp (<_>S1) base [] ], (i = 1) -> <j> comp (<_>S1) (comp (<_>S1) (comp (<_>S1) (comp (<_>S1) base []) [ (j = 1) -> <k> comp (<_>S1) base [] ]) []) [] ]) [ (i = 0) -> <j> comp (<_>S1) (comp (<_>S1) base []) [ (j = 0) -> <k> comp (<_>S1) base [] ], (i = 1) -> <j> comp (<_>S1) (comp (<_>S1) (comp (<_>S1) (comp (<_>S1) (loop {S1} @ -j) []) []) []) [ (j = 0) -> <k> comp (<_>S1) (comp (<_>S1) (comp (<_>S1) base []) []) [] ] ]) [ (i = 0) -> <j> comp (<_>S1) (comp (<_>S1) (comp (<_>S1) base []) []) [ (j = 0) -> <k> comp (<_>S1) (comp (<_>S1) base []) [] ], (i = 1) -> <j> comp (<_>S1) (comp (<_>S1) (comp (<_>S1) (comp (<_>S1) (comp (<_>S1) base [ (j = 0) -> <k> comp (<_>S1) base [ (k = 0) -> <l> base ], (j = 1) -> <k> comp (<_>S1) base [ (k = 0) -> <l> base ] ]) [ (j = 0) -> <k> comp (<_>S1) (comp (<_>S1) base []) [ (k = 0) -> <l> comp (<_>S1) base [] ], (j = 1) -> <k> comp (<_>S1) (comp (<_>S1) base []) [ (k = 0) -> <l> comp (<_>S1) base [] ] ]) [ (j = 0) -> <k> comp (<_>S1) (comp (<_>S1) (comp (<_>S1) base []) []) [ (k = 0) -> <l> comp (<_>S1) (comp (<_>S1) base []) [] ], (j = 1) -> <k> comp (<_>S1) (comp (<_>S1) (comp (<_>S1) base []) []) [ (k = 0) -> <l> comp (<_>S1) (comp (<_>S1) base []) [] ] ]) [ (j = 0) -> <k> comp (<_>S1) (comp (<_>S1) (comp (<_>S1) (comp (<_>S1) base []) []) []) [ (k = 0) -> <l> comp (<_>S1) (comp (<_>S1) (comp (<_>S1) base []) []) [] ], (j = 1) -> <k> comp (<_>S1) (comp (<_>S1) (comp (<_>S1) (comp (<_>S1) base []) []) []) [ (k = 0) -> <l> comp (<_>S1) (comp (<_>S1) (comp (<_>S1) base []) []) [] ] ]) [ (j = 0) -> <k> comp (<_>S1) (comp (<_>S1) (comp (<_>S1) (comp (<_>S1) base []) []) []) [], (j = 1) -> <k> comp (<_>S1) (comp (<_>S1) (comp (<_>S1) (comp (<_>S1) (comp (<_>S1) base []) []) []) []) [ (k = 0) -> <l> comp (<_>S1) (comp (<_>S1) (comp (<_>S1) (comp (<_>S1) base []) []) []) [] ] ] ]) [ (i = 0) -> <j> comp (<_>S1) (comp (<_>S1) (comp (<_>S1) (comp (<_>S1) base []) []) []) [ (j = 0) -> <k> comp (<_>S1) (comp (<_>S1) (comp (<_>S1) base []) []) [] ], (i = 1) -> <j> comp (<_>S1) (comp (<_>S1) (comp (<_>S1) (comp (<_>S1) (comp (<_>S1) base []) []) []) []) [ (j = 1) -> <k> comp (<_>S1) (comp (<_>S1) (comp (<_>S1) (comp (<_>S1) base []) []) []) [] ] ]) [ (i = 0) -> <j> comp (<_>S1) (comp (<_>S1) (comp (<_>S1) (comp (<_>S1) base []) [ (j = 1) -> <k> comp (<_>S1) base [] ]) []) [], (i = 1) -> <j> comp (<_>S1) (comp (<_>S1) (comp (<_>S1) (comp (<_>S1) base []) []) []) [ (j = 1) -> <k> comp (<_>S1) (comp (<_>S1) (comp (<_>S1) base []) []) [] ] ]) [ (i = 0) -> <j> comp (<_>S1) (comp (<_>S1) (comp (<_>S1) base [ (j = 1) -> <k> base ]) []) [], (i = 1) -> <j> comp (<_>S1) (comp (<_>S1) (comp (<_>S1) (comp (<_>S1) (comp (<_>S1) (comp (<_>S1) (comp (<_>S1) (comp (<_>S1) (comp (<_>S1) base [ (j = 1) -> <k> base ]) []) []) [ (j = 0) -> <k> comp (<_>S1) (comp (<_>S1) (comp (<_>S1) base []) []) [ (k = 1) -> <l> comp (<_>S1) (comp (<_>S1) base []) [] ], (j = 1) -> <k> comp (<_>S1) (comp (<_>S1) (comp (<_>S1) base [ (k = 0) -> <l> comp (<_>S1) base [ (l = 0) -> <m> base ], (k = 1) -> <l> comp (<_>S1) base [ (l = 0) -> <m> base ] ]) [ (k = 0) -> <l> comp (<_>S1) (comp (<_>S1) base []) [ (l = 0) -> <m> comp (<_>S1) base [] ], (k = 1) -> <l> comp (<_>S1) (comp (<_>S1) base []) [ (l = 0) -> <m> comp (<_>S1) base [] ] ]) [ (k = 0) -> <l> comp (<_>S1) (comp (<_>S1) base []) [], (k = 1) -> <l> comp (<_>S1) (comp (<_>S1) (comp (<_>S1) base []) []) [ (l = 0) -> <m> comp (<_>S1) (comp (<_>S1) base []) [] ] ] ]) [ (j = 0) -> <k> comp (<_>S1) (comp (<_>S1) base []) [ (k = 1) -> <l> comp (<_>S1) base [] ], (j = 1) -> <k> comp (<_>S1) (comp (<_>S1) (comp (<_>S1) base [ (k = 1) -> <l> base ]) []) [] ]) [ (j = 0) -> <k> comp (<_>S1) (comp (<_>S1) base []) [ (k = 0) -> <l> comp (<_>S1) base [] ], (j = 1) -> <k> comp (<_>S1) (comp (<_>S1) (comp (<_>S1) (loop {S1} @ -k) []) []) [ (k = 0) -> <l> comp (<_>S1) (comp (<_>S1) base []) [] ] ]) [ (j = 0) -> <k> comp (<_>S1) (comp (<_>S1) (comp (<_>S1) base []) []) [ (k = 0) -> <l> comp (<_>S1) (comp (<_>S1) base []) [] ], (j = 1) -> <k> comp (<_>S1) (comp (<_>S1) (comp (<_>S1) (comp (<_>S1) base [ (k = 0) -> <l> comp (<_>S1) base [ (l = 0) -> <m> base ], (k = 1) -> <l> comp (<_>S1) base [ (l = 0) -> <m> base ] ]) [ (k = 0) -> <l> comp (<_>S1) (comp (<_>S1) base []) [ (l = 0) -> <m> comp (<_>S1) base [] ], (k = 1) -> <l> comp (<_>S1) (comp (<_>S1) base []) [ (l = 0) -> <m> comp (<_>S1) base [] ] ]) [ (k = 0) -> <l> comp (<_>S1) (comp (<_>S1) (comp (<_>S1) base []) []) [ (l = 0) -> <m> comp (<_>S1) (comp (<_>S1) base []) [] ], (k = 1) -> <l> comp (<_>S1) (comp (<_>S1) (comp (<_>S1) base []) []) [ (l = 0) -> <m> comp (<_>S1) (comp (<_>S1) base []) [] ] ]) [ (k = 0) -> <l> comp (<_>S1) (comp (<_>S1) (comp (<_>S1) base []) []) [], (k = 1) -> <l> comp (<_>S1) (comp (<_>S1) (comp (<_>S1) (comp (<_>S1) base []) []) []) [ (l = 0) -> <m> comp (<_>S1) (comp (<_>S1) (comp (<_>S1) base []) []) [] ] ] ]) [ (j = 0) -> <k> comp (<_>S1) (comp (<_>S1) (comp (<_>S1) (comp (<_>S1) base []) []) []) [ (k = 0) -> <l> comp (<_>S1) (comp (<_>S1) (comp (<_>S1) base []) []) [] ], (j = 1) -> <k> comp (<_>S1) (comp (<_>S1) (comp (<_>S1) (comp (<_>S1) base []) []) []) [ (k = 1) -> <l> comp (<_>S1) (comp (<_>S1) (comp (<_>S1) base []) []) [] ] ]) [ (j = 0) -> <k> comp (<_>S1) (comp (<_>S1) (comp (<_>S1) (comp (<_>S1) base []) [ (k = 1) -> <l> comp (<_>S1) base [] ]) []) [], (j = 1) -> <k> comp (<_>S1) (comp (<_>S1) (comp (<_>S1) base []) []) [ (k = 1) -> <l> comp (<_>S1) (comp (<_>S1) base []) [] ] ] ]
-- invsone : sone -> sone = subst U (\ (X:U) -> X -> X) <_>S1 sone (<i>s1EqCircle@-i) inv<_>S1
--- cSone : Id U sone sone = <_>sone
+-- cSone : Path U sone sone = <_>sone
-- pt1 : sone =
-- transport cSone
-- (transport cSone
-- (transport cSone (transport cSone north))))))))))
--- lemPt1 : Id sone north pt1 =
+-- lemPt1 : Path sone north pt1 =
-- <i> comp cSone
-- (comp cSone
-- (comp cSone
-- (comp cSone
-- (comp cSone (comp cSone north [(i=0)-><_>north]) [(i=0)-><_>north]) [(i=0)-><_>north]) [(i=0)-><_>north]) [(i=0)-><_>north]) [(i=0)-><_>north]) [(i=0)-><_>north]) [(i=0)-><_>north]) [(i=0)-><_>north]) [(i=0)-><_>north]) [(i=0)-><_>north]
--- transpSone (l:Id sone pt1 pt1) : Id sone north north =
--- compId sone north pt1 north lemPt1 (compId sone pt1 pt1 north l (<i>lemPt1@-i))
+-- transpSone (l:Path sone pt1 pt1) : Path sone north north =
+-- compPath sone north pt1 north lemPt1 (compPath sone pt1 pt1 north l (<i>lemPt1@-i))
-- -- take a lot of time and memory
-- test5 : Z = windingS (transpSone (<i>invsone (loop2S@i)))
{- take a lot of time to type-check
-loopM2 : Id <_>S1 pt0 test1 = mapOnPath S1 S1 invS1 base base loop2
+loopM2 : Path <_>S1 pt0 test1 = mapOnPath S1 S1 invS1 base base loop2
-loopM0 : Id S1 pt0 pt0 = <i>invMult (loop2@i) (loop2@i)
+loopM0 : Path S1 pt0 pt0 = <i>invMult (loop2@i) (loop2@i)
-}
loop @ i -> loop1 @ -i
test (x:S1) : S1 = mult x (invS x)
-testL : Id S1 base base = <i>test (loop{S1} @ i)
+testL : Path S1 base base = <i>test (loop{S1} @ i)
test1 (x:S1) : S1 = mult (invS x) x
-- test0S : Z = winding (<i>mult (loop2@i) (invS (loop2@i)))
-- loop4 : loopS1 = compS1 loop2 loop2
-- test00S : Z = winding (<i>mult (loop4@i) (invS (loop4@i)))
--- lemInv1 : (x:S1) -> Id S1 (mult (invS x) x) base = split
+-- lemInv1 : (x:S1) -> Path S1 (mult (invS x) x) base = split
-- base -> refl S1 base
-- loop @ i -> <j> base -- loop1 @ -i\/i\/j
test2 : loopS1 = <i> mult (loop1@i) (loop1@-i)
--- test1 : IdP (<i>Id S1 (mult (loop1@i) base) (mult (loop1@i) base)) loop1 loop1 = <i j> mult (loop1@i) (loop1@j)
+-- test1 : PathP (<i>Path S1 (mult (loop1@i) base) (mult (loop1@i) base)) loop1 loop1 = <i j> mult (loop1@i) (loop1@j)
test3 : loopS1 = <i> mult (loop1@i) (loop1@i)
--- test3 : Id loopS1 test2 (compS1 loop1 loop1) = refl loopS1 test2
+-- test3 : Path loopS1 test2 (compS1 loop1 loop1) = refl loopS1 test2
--- lemBase (y:S1) : Id S1 (mult base (mult base y)) y =
--- compId S1 (mult base (mult base y)) (mult base y) y (idL (mult base y)) (idL y)
+-- lemBase (y:S1) : Path S1 (mult base (mult base y)) y =
+-- compPath S1 (mult base (mult base y)) (mult base y) y (idL (mult base y)) (idL y)
--- corrInv (x y : S1) : Id S1 (mult x (mult (invS x) y)) y =
--- compId S1 (mult x (mult (invS x) y)) (mult (mult x (invS x)) y) y rem1 rem
--- where rem1 : Id S1 (mult x (mult (invS x) y)) (mult (mult x (invS x)) y) = multAssoc x (invS x) y
+-- corrInv (x y : S1) : Path S1 (mult x (mult (invS x) y)) y =
+-- compPath S1 (mult x (mult (invS x) y)) (mult (mult x (invS x)) y) y rem1 rem
+-- where rem1 : Path S1 (mult x (mult (invS x) y)) (mult (mult x (invS x)) y) = multAssoc x (invS x) y
--- rem : Id S1 (mult (mult x (invS x)) y) y =
--- compId S1 (mult (mult x (invS x)) y) (mult base y) y (<i>mult (lemInv x @i) y) (idL y)
+-- rem : Path S1 (mult (mult x (invS x)) y) y =
+-- compPath S1 (mult (mult x (invS x)) y) (mult base y) y (<i>mult (lemInv x @i) y) (idL y)
--- corrInv1 (x y : S1) : Id S1 (mult (invS x) (mult x y)) y =
--- compId S1 (mult (invS x) (mult x y)) (mult (mult (invS x) x) y) y rem1 rem
--- where rem1 : Id S1 (mult (invS x) (mult x y)) (mult (mult (invS x) x) y) = multAssoc (invS x) x y
+-- corrInv1 (x y : S1) : Path S1 (mult (invS x) (mult x y)) y =
+-- compPath S1 (mult (invS x) (mult x y)) (mult (mult (invS x) x) y) y rem1 rem
+-- where rem1 : Path S1 (mult (invS x) (mult x y)) (mult (mult (invS x) x) y) = multAssoc (invS x) x y
--- rem : Id S1 (mult (mult (invS x) x) y) y =
--- compId S1 (mult (mult (invS x) x) y) (mult base y) y (<i>mult (lemInv1 x@i) y) (idL y)
+-- rem : Path S1 (mult (mult (invS x) x) y) y =
+-- compPath S1 (mult (mult (invS x) x) y) (mult base y) y (<i>mult (lemInv1 x@i) y) (idL y)
--- eqS1 (x:S1) : Id U S1 S1 = isoId S1 S1 (mult x) (mult (invS x)) (corrInv x) (corrInv1 x)
+-- eqS1 (x:S1) : Path U S1 S1 = isoPath S1 S1 (mult x) (mult (invS x)) (corrInv x) (corrInv1 x)
zero -> m
suc n -> suc (add m n)
-add_zero : (n : nat) -> Id nat (add zero n) n = split
+add_zero : (n : nat) -> Path nat (add zero n) n = split
zero -> <i> zero
suc n -> <i> suc (add_zero n @ i)
-add_suc (a:nat) : (n : nat) -> Id nat (add (suc a) n) (suc (add a n)) = split
+add_suc (a:nat) : (n : nat) -> Path nat (add (suc a) n) (suc (add a n)) = split
zero -> <i> suc a
suc m -> <i> suc (add_suc a m @ i)
-add_comm (a : nat) : (n : nat) -> Id nat (add a n) (add n a) = split
+add_comm (a : nat) : (n : nat) -> Path nat (add a n) (add n a) = split
zero -> <i> add_zero a @ -i
suc m -> <i> comp (<_> nat) (suc (add_comm a m @ i))
[ (i = 0) -> <j> suc (add a m)
, (i = 1) -> <j> add_suc m a @ -j ]
-assocAdd (a b:nat) : (c:nat) -> Id nat (add a (add b c)) (add (add a b) c) = split
+assocAdd (a b:nat) : (c:nat) -> Path nat (add a (add b c)) (add (add a b) c) = split
zero -> <i>add a b
suc c1 -> <i>suc (assocAdd a b c1@i)
zero -> \(x : nat) -> x
suc n -> \(x : nat) -> suc (add' n x)
-sucInj (n m : nat) (p : Id nat (suc n) (suc m)) : Id nat n m =
+sucInj (n m : nat) (p : Path nat (suc n) (suc m)) : Path nat n m =
<i> pred (p @ i)
-add_comm3 (a b c : nat) : Id nat (add a (add b c)) (add c (add b a)) =
- let rem : Id nat (add a (add b c)) (add a (add c b)) = <i> add a (add_comm b c @ i)
- rem1 : Id nat (add a (add c b)) (add (add c b) a) = add_comm a (add c b)
- rem2 : Id nat (add (add c b) a) (add c (add b a)) = <i> assocAdd c b a @ -i
+add_comm3 (a b c : nat) : Path nat (add a (add b c)) (add c (add b a)) =
+ let rem : Path nat (add a (add b c)) (add a (add c b)) = <i> add a (add_comm b c @ i)
+ rem1 : Path nat (add a (add c b)) (add (add c b) a) = add_comm a (add c b)
+ rem2 : Path nat (add (add c b) a) (add c (add b a)) = <i> assocAdd c b a @ -i
in <i> comp (<_> nat) (rem1 @ i) [ (i = 0) -> <j> rem @ -j, (i = 1) -> rem2 ]
-natcancelr (a b : nat) : (x : nat) -> Id nat (add a x) (add b x) -> Id nat a b = split
- zero -> \(h : Id nat a b) -> h
- suc x' -> \(h : Id nat (suc (add a x')) (suc (add b x'))) ->
+natcancelr (a b : nat) : (x : nat) -> Path nat (add a x) (add b x) -> Path nat a b = split
+ zero -> \(h : Path nat a b) -> h
+ suc x' -> \(h : Path nat (suc (add a x')) (suc (add b x'))) ->
natcancelr a b x' (sucInj (add a x') (add b x') h)
idnat : nat -> nat = split
zero -> zero
suc n -> suc (idnat n)
-test : Id (nat -> nat) idnat (idfun nat) = funExt nat (\(_ : nat) -> nat) idnat (idfun nat) rem
+test : Path (nat -> nat) idnat (idfun nat) = funExt nat (\(_ : nat) -> nat) idnat (idfun nat) rem
where
- rem : (n : nat) -> Id nat (idnat n) n = split
+ rem : (n : nat) -> Path nat (idnat n) n = split
zero -> refl nat zero
suc n -> mapOnPath nat nat (\(x : nat) -> suc x) (idnat n) n (rem n)
zero -> a0
suc n -> aS n
-znots (n : nat) : neg (Id nat zero (suc n)) =
- \ (h:Id nat zero (suc n)) -> subst nat (caseNat U nat N0) zero (suc n) h zero
+znots (n : nat) : neg (Path nat zero (suc n)) =
+ \ (h:Path nat zero (suc n)) -> subst nat (caseNat U nat N0) zero (suc n) h zero
-snotz (n : nat) : neg (Id nat (suc n) zero) =
- \ (h:Id nat (suc n) zero) -> znots n (inv nat (suc n) zero h)
+snotz (n : nat) : neg (Path nat (suc n) zero) =
+ \ (h:Path nat (suc n) zero) -> znots n (inv nat (suc n) zero h)
-natDec : (n m:nat) -> dec (Id nat n m) = split
- zero -> caseDNat (\ (m:nat) -> dec (Id nat zero m)) (inl (refl nat zero)) (\ (m:nat) -> inr (znots m))
- suc n -> caseDNat (\ (m:nat) -> dec (Id nat (suc n) m)) (inr (snotz n))
- (\ (m:nat) -> decEqCong (Id nat n m) (Id nat (suc n) (suc m)) (\ (p:Id nat n m) -> <i> suc (p @ i))
+natDec : (n m:nat) -> dec (Path nat n m) = split
+ zero -> caseDNat (\ (m:nat) -> dec (Path nat zero m)) (inl (refl nat zero)) (\ (m:nat) -> inr (znots m))
+ suc n -> caseDNat (\ (m:nat) -> dec (Path nat (suc n) m)) (inr (snotz n))
+ (\ (m:nat) -> decEqCong (Path nat n m) (Path nat (suc n) (suc m)) (\ (p:Path nat n m) -> <i> suc (p @ i))
(sucInj n m) (natDec n m))
natSet : set nat = hedberg nat natDec
zero -> Unit
succ z -> O2 n z
lim f -> (p:nat) -> O2 n (f p)
- lim2 f -> (x:ord) -> and (O2 n (f x)) (Id ord (G2 (f x) n) (G2 (f (inj0 (G1 x n))) n))
+ lim2 f -> (x:ord) -> and (O2 n (f x)) (Path ord (G2 (f x) n) (G2 (f (inj0 (G1 x n))) n))
inj12 : ord -> ord2 = split
zero -> zero
lim2 f -> \ (x:ord) -> H2 (f x) x
collapsing (n:nat) :
- (x:ord2) (y:ord) -> O2 n x -> Id nat (G1 (H2 x y) n) (H1 (G2 x n) (G1 y n)) = split
+ (x:ord2) (y:ord) -> O2 n x -> Path nat (G1 (H2 x y) n) (H1 (G2 x n) (G1 y n)) = split
zero -> \ (y:ord) (h:O2 n zero) -> <i>G1 y n
succ z -> \ (y:ord) (h:O2 n (succ z)) -> collapsing n z (succ y) h
lim f -> \ (y:ord) (h:O2 n (lim f)) -> collapsing n (f n) y (h n)
lim2 f -> \ (y:ord) (h:O2 n (lim2 f)) ->
let
- rem : Id ord (G2 (f y) n) (G2 (f (inj0 (G1 y n))) n) = (h y).2
- rem1 : Id nat (G1 (H2 (f y) y) n) (H1 (G2 (f y) n) (G1 y n)) = collapsing n (f y) y (h y).1
- in comp (<i>Id nat (G1 (H2 (f y) y) n) (H1 (rem@i) (G1 y n))) rem1 []
+ rem : Path ord (G2 (f y) n) (G2 (f (inj0 (G1 y n))) n) = (h y).2
+ rem1 : Path nat (G1 (H2 (f y) y) n) (H1 (G2 (f y) n) (G1 y n)) = collapsing n (f y) y (h y).1
+ in comp (<i>Path nat (G1 (H2 (f y) y) n) (H1 (rem@i) (G1 y n))) rem1 []
-- an application
zero -> tt
succ z -> rem z
lim f -> \ (p:nat) -> rem (f p)
- rem1 : (x:ord) -> Id ord (G2 (inj12 x) n) (G2 (inj12 (inj0 (G1 x n))) n) = split
+ rem1 : (x:ord) -> Path ord (G2 (inj12 x) n) (G2 (inj12 (inj0 (G1 x n))) n) = split
zero -> <i>zero
succ z -> <i>succ ((rem1 z)@i)
lim f -> rem1 (f n)
-corr1 (n:nat) : Id nat (G1 (H2 omega1 omega) n) (H1 (G2 omega1 n) (G1 omega n)) =
+corr1 (n:nat) : Path nat (G1 (H2 omega1 omega) n) (H1 (G2 omega1 n) (G1 omega n)) =
collapsing n omega1 omega (lemOmega1 n)
-lem : (n p:nat) -> Id nat (G1 (inj0 n) p) n = split
+lem : (n p:nat) -> Path nat (G1 (inj0 n) p) n = split
zero -> \ (p:nat) -> <i>zero
succ q -> \ (p:nat) -> <i> succ (lem q p@i)
-lem1 (n:nat) : Id nat (G1 omega n) n = lem n n
+lem1 (n:nat) : Path nat (G1 omega n) n = lem n n
-lem2 : (n p:nat) -> Id ord (G2 (inj12 (inj0 n)) p) (inj0 n) = split
+lem2 : (n p:nat) -> Path ord (G2 (inj12 (inj0 n)) p) (inj0 n) = split
zero -> \ (p:nat) -> <i>inj0 zero
succ q -> \ (p:nat) -> <i> succ (lem2 q p@i)
test (n:nat) : ord = G2 omega1 n
-lem3 (n:nat) : Id ord (G2 (inj12 (inj0 n)) n) (inj0 n) = lem2 n n
+lem3 (n:nat) : Path ord (G2 (inj12 (inj0 n)) n) (inj0 n) = lem2 n n
-lem4 (n:nat) : Id nat (H1 (G2 omega1 n) n) (H1 omega n) =
+lem4 (n:nat) : Path nat (H1 (G2 omega1 n) n) (H1 omega n) =
<i>H1 (lem3 n@i) n
-- the G1 and H1 hierarchy coincides: G1 (H2 omega1 omega) and H1 omega are the same function
-corr2 : Id (nat -> nat) (G1 (H2 omega1 omega)) (H1 omega) =
+corr2 : Path (nat -> nat) (G1 (H2 omega1 omega)) (H1 omega) =
<i>\ (n:nat) -> comp (<_>nat) (H1 (G2 omega1 n) ((lem1 n)@i)) [(i=0) -> <j>corr1 n@-j,(i=1) -> lem4 n]
pi (A:U) (P:A->U) : U = (x:A) -> P x
-idPi (A:U) (B:A->U) (f g : pi A B) : Id U (Id (pi A B) f g) ((x:A) -> Id (B x) (f x) (g x)) =
- isoId (Id (pi A B) f g) ((x:A) -> Id (B x) (f x) (g x)) F G S T
- where T0 : U = Id (pi A B) f g
- T1 : U = (x:A) -> Id (B x) (f x) (g x)
+idPi (A:U) (B:A->U) (f g : pi A B) : Path U (Path (pi A B) f g) ((x:A) -> Path (B x) (f x) (g x)) =
+ isoPath (Path (pi A B) f g) ((x:A) -> Path (B x) (f x) (g x)) F G S T
+ where T0 : U = Path (pi A B) f g
+ T1 : U = (x:A) -> Path (B x) (f x) (g x)
F (p:T0) : T1 = \ (x:A) -> <i>p@i x
G (p:T1) : T0 = <i>\ (x:A) -> p x @ i
- S (p:T1) : Id T1 (F (G p)) p = refl T1 p
- T (p:T0) : Id T0 (G (F p)) p = refl T0 p
+ S (p:T1) : Path T1 (F (G p)) p = refl T1 p
+ T (p:T0) : Path T0 (G (F p)) p = refl T0 p
-setPi (A:U) (B:A -> U) (h:(x:A) -> set (B x)) (f g:pi A B) : prop (Id (pi A B) f g) =
+setPi (A:U) (B:A -> U) (h:(x:A) -> set (B x)) (f g:pi A B) : prop (Path (pi A B) f g) =
rem
where
- T : U = (x:A) -> Id (B x) (f x) (g x)
+ T : U = (x:A) -> Path (B x) (f x) (g x)
rem1 : prop T = \ (p q : T) -> <i> \ (x:A) -> h x (f x) (g x) (p x) (q x)@i
- rem : prop (Id (pi A B) f g) =
- subst U prop T (Id (pi A B) f g) (<i>idPi A B f g@-i) rem1
+ rem : prop (Path (pi A B) f g) =
+ subst U prop T (Path (pi A B) f g) (<i>idPi A B f g@-i) rem1
-groupoidPi (A:U) (B:A -> U) (h:(x:A) -> groupoid (B x)) (f g:pi A B) : set (Id (pi A B) f g) =
- subst U set T (Id (pi A B) f g) (<i>idPi A B f g@-i) rem1
+groupoidPi (A:U) (B:A -> U) (h:(x:A) -> groupoid (B x)) (f g:pi A B) : set (Path (pi A B) f g) =
+ subst U set T (Path (pi A B) f g) (<i>idPi A B f g@-i) rem1
where
- T : U = (x:A) -> Id (B x) (f x) (g x)
- rem1 : set T = setPi A (\ (x:A) -> Id (B x) (f x) (g x)) (\ (x:A) -> h x (f x) (g x))
+ T : U = (x:A) -> Path (B x) (f x) (g x)
+ rem1 : set T = setPi A (\ (x:A) -> Path (B x) (f x) (g x)) (\ (x:A) -> h x (f x) (g x))
-- Identity types
-Id (A : U) (a0 a1 : A) : U = IdP (<i> A) a0 a1
+Path (A : U) (a0 a1 : A) : U = PathP (<i> A) a0 a1
-refl (A : U) (a : A) : Id A a a = <i> a
+refl (A : U) (a : A) : Path A a a = <i> a
-testEta (A : U) (a b : A) (p : Id A a b) : Id (Id A a b) p p = refl (Id A a b) (<i> p @ i)
+testEta (A : U) (a b : A) (p : Path A a b) : Path (Path A a b) p p = refl (Path A a b) (<i> p @ i)
mapOnPath (A B : U) (f : A -> B) (a b : A)
- (p : Id A a b) : Id B (f a) (f b) = <i> f (p @ i)
+ (p : Path A a b) : Path B (f a) (f b) = <i> f (p @ i)
funExt (A : U) (B : A -> U) (f g : (x : A) -> B x)
- (p : (x : A) -> Id (B x) (f x) (g x)) :
- Id ((y : A) -> B y) f g = <i> \(a : A) -> (p a) @ i
+ (p : (x : A) -> Path (B x) (f x) (g x)) :
+ Path ((y : A) -> B y) f g = <i> \(a : A) -> (p a) @ i
-- Transport can be defined using comp
-trans (A B : U) (p : Id U A B) (a : A) : B = comp p a []
+trans (A B : U) (p : Path U A B) (a : A) : B = comp p a []
-- subst can be defined using trans:
-substTrans (A : U) (P : A -> U) (a b : A) (p : Id A a b) (e : P a) : P b =
+substTrans (A : U) (P : A -> U) (a b : A) (p : Path A a b) (e : P a) : P b =
trans (P a) (P b) (mapOnPath A U P a b p) e
-subst (A : U) (P : A -> U) (a b : A) (p : Id A a b) (e : P a) : P b =
+subst (A : U) (P : A -> U) (a b : A) (p : Path A a b) (e : P a) : P b =
transport (mapOnPath A U P a b p) e
substEq (A : U) (P : A -> U) (a : A) (e : P a)
- : Id (P a) e (subst A P a a (refl A a) e) =
+ : Path (P a) e (subst A P a a (refl A a) e) =
fill (<i> P a) e []
-substInv (A : U) (P : A -> U) (a b : A) (p : Id A a b) : P b -> P a =
+substInv (A : U) (P : A -> U) (a b : A) (p : Path A a b) : P b -> P a =
subst A P b a (<i> p @ -i)
-singl (A : U) (a : A) : U = (x : A) * Id A a x
+singl (A : U) (a : A) : U = (x : A) * Path A a x
-contrSingl (A : U) (a b : A) (p : Id A a b) :
- Id (singl A a) (a,refl A a) (b,p) = <i> (p @ i,<j> p @ i/\j)
+contrSingl (A : U) (a b : A) (p : Path A a b) :
+ Path (singl A a) (a,refl A a) (b,p) = <i> (p @ i,<j> p @ i/\j)
-J (A : U) (a : A) (C : (x : A) -> Id A a x -> U)
- (d : C a (refl A a)) (x : A) (p : Id A a x) : C x p =
+J (A : U) (a : A) (C : (x : A) -> Path A a x -> U)
+ (d : C a (refl A a)) (x : A) (p : Path A a x) : C x p =
subst (singl A a) T (a, refl A a) (x, p) (contrSingl A a x p) d
where T (z : singl A a) : U = C (z.1) (z.2)
-JEq (A : U) (a : A) (C : (x : A) -> Id A a x -> U) (d : C a (refl A a))
- : Id (C a (refl A a)) d (J A a C d a (refl A a)) =
+JEq (A : U) (a : A) (C : (x : A) -> Path A a x -> U) (d : C a (refl A a))
+ : Path (C a (refl A a)) d (J A a C d a (refl A a)) =
substEq (singl A a) T (a, refl A a) d
where T (z : singl A a) : U = C (z.1) (z.2)
-inv (A : U) (a b : A) (p : Id A a b) : Id A b a = <i> p @ -i
+inv (A : U) (a b : A) (p : Path A a b) : Path A b a = <i> p @ -i
-compId (A : U) (a b c : A) (p : Id A a b) (q : Id A b c) : Id A a c =
+compPath (A : U) (a b c : A) (p : Path A a b) (q : Path A b c) : Path A a c =
<i> comp (<j>A) (p @ i) [ (i = 1) -> q, (i=0) -> <j> a ]
-compId' (A : U) (a b c : A) (p : Id A a b) (q : Id A b c) : Id A a c =
- subst A (Id A a) b c q p
+compPath' (A : U) (a b c : A) (p : Path A a b) (q : Path A b c) : Path A a c =
+ subst A (Path A a) b c q p
-compId'' (A : U) (a b : A) (p : Id A a b) : (c : A) -> (q : Id A b c) -> Id A a c =
- J A a ( \ (b : A) (p : Id A a b) -> (c : A) -> (q : Id A b c) -> Id A a c) rem b p
- where rem (c : A) (p : Id A a c) : Id A a c = p
+compPath'' (A : U) (a b : A) (p : Path A a b) : (c : A) -> (q : Path A b c) -> Path A a c =
+ J A a ( \ (b : A) (p : Path A a b) -> (c : A) -> (q : Path A b c) -> Path A a c) rem b p
+ where rem (c : A) (p : Path A a c) : Path A a c = p
compUp (A : U) (a a' b b' : A)
- (p : Id A a a') (q : Id A b b') (r : Id A a b) : Id A a' b' =
+ (p : Path A a a') (q : Path A b b') (r : Path A a b) : Path A a' b' =
<i> comp (<j>A) (r @ i) [(i = 0) -> p, (i = 1) -> q]
compDown (A : U) (a a' b b' : A)
- (p : Id A a a') (q: Id A b b') : Id A a' b' -> Id A a b =
+ (p : Path A a a') (q: Path A b b') : Path A a' b' -> Path A a b =
compUp A a' a b' b (inv A a a' p) (inv A b b' q)
-lemCompInv (A:U) (a b c:A) (p:Id A a b) (q:Id A b c)
- : Id (Id A a b) (compId A a c b (compId A a b c p q) (inv A b c q)) p =
+lemCompInv (A:U) (a b c:A) (p:Path A a b) (q:Path A b c)
+ : Path (Path A a b) (compPath A a c b (compPath A a b c p q) (inv A b c q)) p =
<j i> comp (<k>A)
((fill (<k>A) (p @ i) [(i=0) -> <k>a, (i=1) -> q]) @ -j)
[ (i=0) -> <k> a
, (i=1) -> <k> q @ - (j \/ k)
- , (j=0) -> fill (<k>A) ((compId A a b c p q @ i))
+ , (j=0) -> fill (<k>A) ((compPath A a b c p q @ i))
[(i=0) -> <k>a, (i=1) -> <k> q @ -k ]
, (j=1) -> <k> p @ i
]
-lemInv (A:U) (a b:A) (p:Id A a b) : Id (Id A b b) (compId A b a b (inv A a b p) p) (refl A b) =
+lemInv (A:U) (a b:A) (p:Path A a b) : Path (Path A b b) (compPath A b a b (inv A a b p) p) (refl A b) =
<j i> comp (<k>A) (p @ (-i \/ j)) [(i=0) -> <l>b, (j=1) -> <l>b, (i=1) -> <k> p @ (j \/ k)]
-test0 (A : U) (a b : A) (p : Id A a b) : Id A a a = refl A (p @ 0)
-test1 (A : U) (a b : A) (p : Id A a b) : Id A b b = refl A (p @ 1)
+test0 (A : U) (a b : A) (p : Path A a b) : Path A a a = refl A (p @ 0)
+test1 (A : U) (a b : A) (p : Path A a b) : Path A b b = refl A (p @ 1)
--- compEmpty (A : U) (a b : A) (p : Id A a b) : Id A a b =
+-- compEmpty (A : U) (a b : A) (p : Path A a b) : Path A a b =
-- <i> comp A (p @ i) [ ]
-kan (A : U) (a b c d : A) (p : Id A a b) (q : Id A a c)
- (r : Id A b d) : Id A c d =
+kan (A : U) (a b c d : A) (p : Path A a b) (q : Path A a c)
+ (r : Path A b d) : Path A c d =
<i> comp (<j>A) (p @ i) [ (i = 0) -> q, (i = 1) -> r ]
-lemSimpl (A : U) (a b c : A) (p : Id A a b) (q q' : Id A b c)
- (s : Id (Id A a c) (compId A a b c p q) (compId A a b c p q')) :
- Id (Id A b c) q q' =
+lemSimpl (A : U) (a b c : A) (p : Path A a b) (q q' : Path A b c)
+ (s : Path (Path A a c) (compPath A a b c p q) (compPath A a b c p q')) :
+ Path (Path A b c) q q' =
<j k> comp (<i> A) a
[ (j = 0) -> <i> comp (<l> A) (p @ i)
[ (k = 0) -> <l> p @ i
, (k = 0) -> p
, (k = 1) -> s @ j ]
-IdPathTest1 (A : U) (a b : A) (p : Id A a b) :
- Id (Id A a b) p (<i> comp (<j> A) (p @ i) [(i=0) -> <j> a,(i=1) -> <j> b]) =
+PathPathTest1 (A : U) (a b : A) (p : Path A a b) :
+ Path (Path A a b) p (<i> comp (<j> A) (p @ i) [(i=0) -> <j> a,(i=1) -> <j> b]) =
<j i> fill (<k> A) (p @ i) [(i=0) -> <k> a,(i=1) -> <k> b] @ j
idfun (A : U) (a : A) : A = a
-- b0 -----> b1
-- v
Square (A : U) (a0 a1 b0 b1 : A)
- (u : Id A a0 a1) (v : Id A b0 b1)
- (r0 : Id A a0 b0) (r1 : Id A a1 b1) : U
- = IdP (<i> (IdP (<j> A) (u @ i) (v @ i))) r0 r1
+ (u : Path A a0 a1) (v : Path A b0 b1)
+ (r0 : Path A a0 b0) (r1 : Path A a1 b1) : U
+ = PathP (<i> (PathP (<j> A) (u @ i) (v @ i))) r0 r1
-constSquare (A : U) (a : A) (p : Id A a a) : Square A a a a a p p p p =
+constSquare (A : U) (a : A) (p : Path A a a) : Square A a a a a p p p p =
<i j> comp (<k>A) a
[(i = 0) -> <k> p @ (j \/ - k),
(i = 1) -> <k> p @ (j /\ k),
(j = 0) -> <k> p @ (i \/ - k),
(j = 1) -> <k> p @ (i /\ k)]
-prop (A : U) : U = (a b : A) -> Id A a b
-set (A : U) : U = (a b : A) -> prop (Id A a b)
-groupoid (A : U) : U = (a b : A) -> set (Id A a b)
+prop (A : U) : U = (a b : A) -> Path A a b
+set (A : U) : U = (a b : A) -> prop (Path A a b)
+groupoid (A : U) : U = (a b : A) -> set (Path A a b)
-- the collection of all sets
SET : U = (X:U) * set X
propSet (A : U) (h : prop A) : set A =
- \(a b : A) (p q : Id A a b) ->
+ \(a b : A) (p q : Path A a b) ->
<j i> comp (<k>A) a [ (i=0) -> h a a
, (i=1) -> h a b
, (j=0) -> h a (p @ i)
setIsProp (A : U) : prop (set A) =
\(f g : set A) -> <i> \(a b :A) ->
- propIsProp (Id A a b) (f a b) (g a b) @ i
+ propIsProp (Path A a b) (f a b) (g a b) @ i
-IdS (A : U) (P : A -> U) (a0 a1 : A)
- (p : Id A a0 a1) (u0 : P a0) (u1 : P a1) : U =
- IdP (<i> P (p @ i)) u0 u1
+PathS (A : U) (P : A -> U) (a0 a1 : A)
+ (p : Path A a0 a1) (u0 : P a0) (u1 : P a1) : U =
+ PathP (<i> P (p @ i)) u0 u1
lemProp (A : U) (h : A -> prop A) : prop A =
\(a : A) -> h a a
propPi (A : U) (B : A -> U) (h : (x : A) -> prop (B x))
- (f0 f1 : (x : A) -> B x) : Id ((x : A) -> B x) f0 f1
+ (f0 f1 : (x : A) -> B x) : Path ((x : A) -> B x) f0 f1
= <i> \ (x:A) -> (h x (f0 x) (f1 x)) @ i
lemPropF (A : U) (P : A -> U) (pP : (x : A) -> prop (P x)) (a0 a1 :A)
- (p : Id A a0 a1) (b0 : P a0) (b1 : P a1) : IdP (<i>P (p@i)) b0 b1 =
+ (p : Path A a0 a1) (b0 : P a0) (b1 : P a1) : PathP (<i>P (p@i)) b0 b1 =
<i>pP (p@i) (comp (<j>P (p@i/\j)) b0 [(i=0) -> <_>b0]) (comp (<j>P (p@i\/-j)) b1 [(i=1) -> <_>b1])@i
-- other proof
-- lemPropF (A : U) (P : A -> U) (pP : (x : A) -> prop (P x)) (a :A) :
--- (a1 : A) (p : Id A a a1) (b0 : P a) (b1 : P a1) -> IdP (<i>P (p@i)) b0 b1 =
--- J A a (\ (a1 : A) (p : Id A a a1) ->
--- (b0 : P a) (b1 : P a1) -> IdP (<i>P (p@i)) b0 b1)
+-- (a1 : A) (p : Path A a a1) (b0 : P a) (b1 : P a1) -> PathP (<i>P (p@i)) b0 b1 =
+-- J A a (\ (a1 : A) (p : Path A a a1) ->
+-- (b0 : P a) (b1 : P a1) -> PathP (<i>P (p@i)) b0 b1)
-- rem
--- where rem : (b0 b1:P a) -> Id (P a) b0 b1 = pP a
+-- where rem : (b0 b1:P a) -> Path (P a) b0 b1 = pP a
Sigma (A : U) (B : A -> U) : U = (x : A) * B x
lemSig (A : U) (B : A -> U) (pB : (x : A) -> prop (B x))
- (u v : (x:A) * B x) (p : Id A u.1 v.1) :
- Id ((x:A) * B x) u v =
+ (u v : (x:A) * B x) (p : Path A u.1 v.1) :
+ Path ((x:A) * B x) u v =
<i> (p@i,(lemPropF A B pB u.1 v.1 p u.2 v.2)@i)
propSig (A : U) (B : A -> U) (pA : prop A)
(pB : (x : A) -> prop (B x)) (t u : (x:A) * B x) :
- Id ((x:A) * B x) t u =
+ Path ((x:A) * B x) t u =
lemSig A B pB t u (pA t.1 u.1)
-isContr (A : U) : U = (x : A) * ((y : A) -> Id A x y)
+isContr (A : U) : U = (x : A) * ((y : A) -> Path A x y)
propIsContr (A : U) : prop (isContr A) = lemProp (isContr A) rem
where
rem (t : isContr A) : prop (isContr A) = propSig A T pA pB
where
- T (x : A) : U = (y : A) -> Id A x y
- pA (x y : A) : Id A x y = compId A x t.1 y (<i> t.2 x @ -i) (t.2 y)
+ T (x : A) : U = (y : A) -> Path A x y
+ pA (x y : A) : Path A x y = compPath A x t.1 y (<i> t.2 x @ -i) (t.2 y)
pB (x : A) : prop (T x) =
- propPi A (\ (y : A) -> Id A x y) (propSet A pA x)
+ propPi A (\ (y : A) -> Path A x y) (propSet A pA x)
-- Alternative proof:
--- propIsContr (A:U) (z0 z1:isContr A) : Id (isContr A) z0 z1 =
+-- propIsContr (A:U) (z0 z1:isContr A) : Path (isContr A) z0 z1 =
-- <j>(p0 a1@j,
-- \ (x:A) ->
-- <i>comp (<_>A) (lem1 x@i@j) [ (i=0) -> <k>p0 a1@j, (i=1) -> <k>p0 x@(j\/k),
-- (j=0) -> <k>p0 x@(i/\k), (j=1) -> <k>p1 x@i ])
-- where
-- a0 : A = z0.1
--- p0 : (x:A) -> Id A a0 x = z0.2
+-- p0 : (x:A) -> Path A a0 x = z0.2
-- a1 : A = z1.1
--- p1 : (x:A) -> Id A a1 x = z1.2
--- lem1 (x:A) : IdP (<i>Id A a0 (p1 x@i)) (p0 a1) (p0 x) = <i j> p0 (p1 x@i) @ j
+-- p1 : (x:A) -> Path A a1 x = z1.2
+-- lem1 (x:A) : PathP (<i>Path A a0 (p1 x@i)) (p0 a1) (p0 x) = <i j> p0 (p1 x@i) @ j
-- Basic data types
data Unit = tt
propUnit : prop Unit = split
- tt -> split@((x:Unit) -> Id Unit tt x) with
+ tt -> split@((x:Unit) -> Path Unit tt x) with
tt -> <i> tt
setUnit : set Unit = propSet Unit propUnit
| inr (b : B)
propOr (A B : U) (hA : prop A) (hB : prop B) (h : A -> neg B) : prop (or A B) = split
- inl a' -> split@((b : or A B) -> Id (or A B) (inl a') b) with
+ inl a' -> split@((b : or A B) -> Path (or A B) (inl a') b) with
inl b' -> <i> inl (hA a' b' @ i)
- inr b' -> efq (Id (or A B) (inl a') (inr b')) (h a' b')
- inr a' -> split@((b : or A B) -> Id (or A B) (inr a') b) with
- inl b' -> efq (Id (or A B) (inr a') (inl b')) (h b' a')
+ inr b' -> efq (Path (or A B) (inl a') (inr b')) (h a' b')
+ inr a' -> split@((b : or A B) -> Path (or A B) (inr a') b) with
+ inl b' -> efq (Path (or A B) (inr a') (inl b')) (h b' a')
inr b' -> <i> inr (hB a' b' @ i)
stable (A:U) : U = neg (neg A) -> A
-const (A : U) (f : A -> A) : U = (x y : A) -> Id A (f x) (f y)
+const (A : U) (f : A -> A) : U = (x y : A) -> Path A (f x) (f y)
exConst (A : U) : U = (f:A -> A) * const A f
-propN0 : prop N0 = \ (x y:N0) -> efq (Id N0 x y) x
+propN0 : prop N0 = \ (x y:N0) -> efq (Path N0 x y) x
propNeg (A:U) : prop (neg A) = \ (f g:neg A) -> <i>\(x:A) -> (propN0 (f x) (g x))@i
decConst (A : U) : dec A -> exConst A = split
inl a -> (\ (x:A) -> a, \ (x y:A) -> refl A a)
- inr h -> (\ (x:A) -> x, \ (x y:A) -> efq (Id A x y) (h x))
+ inr h -> (\ (x:A) -> x, \ (x y:A) -> efq (Path A x y) (h x))
stableConst (A : U) (sA: stable A) : exConst A =
(\ (x:A) -> sA (dNeg A x),\ (x y:A) -> <i>sA (propNeg (neg A) (dNeg A x) (dNeg A y) @ i))
-discrete (A : U) : U = (a b : A) -> dec (Id A a b)
+discrete (A : U) : U = (a b : A) -> dec (Path A a b)
injective (A B : U) (f : A -> B) : U =
- (a0 a1 : A) -> Id B (f a0) (f a1) -> Id A a0 a1
+ (a0 a1 : A) -> Path B (f a0) (f a1) -> Path A a0 a1
and (A B : U) : U = (_ : A) * B
import equiv
lemProp (A B : U) (f : A -> B) (g : B -> A)
- (s : (y:B) -> Id B (f (g y)) y)
- (t : (x:A) -> Id A (g (f x)) x) (pA:prop A) : prop B =
- \ (b0 b1:B) -> subst U prop A B (isoId A B f g s t) pA b0 b1
+ (s : (y:B) -> Path B (f (g y)) y)
+ (t : (x:A) -> Path A (g (f x)) x) (pA:prop A) : prop B =
+ \ (b0 b1:B) -> subst U prop A B (isoPath A B f g s t) pA b0 b1
{- normal form
\(A B : U) -> \(f : A -> B) -> \(g : B -> A) ->
- \(s : (y : B) -> IdP (<!0> B) (f (g y)) y) ->
- \(t : (x : A) -> IdP (<!0> A) (g (f x)) x) ->
- \(pA : (a : A) -> (b : A) -> IdP (<!0> A) a b) ->
+ \(s : (y : B) -> PathP (<!0> B) (f (g y)) y) ->
+ \(t : (x : A) -> PathP (<!0> A) (g (f x)) x) ->
+ \(pA : (a : A) -> (b : A) -> PathP (<!0> A) a b) ->
\(b0 b1 : B) ->
<!1> comp B (f (pA (g b0) (g b1) @ !1))
[ (!1 = 0) -> <!2> comp B b0
[ (!3 = 0) -> <!4> s b1 @ !4, (!3 = 1) -> <!4> s (f (g b1)) @ !4 ] ] ]
\(A B : U) -> \(f : A -> B) -> \(g : B -> A) ->
- \(s : (y : B) -> IdP (<!0> B) (f (g y)) y) ->
- \(t : (x : A) -> IdP (<!0> A) (g (f x)) x) ->
- \(pA : (a b : A) -> IdP (<!0> A) a b) ->
+ \(s : (y : B) -> PathP (<!0> B) (f (g y)) y) ->
+ \(t : (x : A) -> PathP (<!0> A) (g (f x)) x) ->
+ \(pA : (a b : A) -> PathP (<!0> A) a b) ->
\(b0 b1 : B) ->
<!1> comp B (f (pA (g b0) (g b1) @ !1)) [ (!1 = 0) -> <!2> comp B b0 [ (!2 = 0) -> <!3> comp B (f (comp A (g b0) [ (!3 = 1) -> <!4> t (g b0) @ -!4 ])) [ (!3 = 0) -> <!4> s b0 @ !4, (!3 = 1) -> <!4> s (f (g b0)) @ !4 ] ], (!1 = 1) -> <!2> comp B b1 [ (!2 = 0) -> <!3> comp B (f (comp A (g b1) [ (!3 = 1) -> <!4> t (g b1) @ -!4 ])) [ (!3 = 0) -> <!4> s b1 @ !4, (!3 = 1) -> <!4> s (f (g b1)) @ !4 ] ] ]
-}
lemSet (A B : U) (f : A -> B) (g : B -> A)
- (s : (y:B) -> Id B (f (g y)) y)
- (t : (x:A) -> Id A (g (f x)) x) (sA:set A) : set B =
- \ (b0 b1:B) (p q : Id B b0 b1) -> subst U set A B (isoId A B f g s t) sA b0 b1 p q
+ (s : (y:B) -> Path B (f (g y)) y)
+ (t : (x:A) -> Path A (g (f x)) x) (sA:set A) : set B =
+ \ (b0 b1:B) (p q : Path B b0 b1) -> subst U set A B (isoPath A B f g s t) sA b0 b1 p q
{- normal form
\(A B : U) -> \(f : A -> B) -> \(g : B -> A) ->
- \(s : (y : B) -> IdP (<!0> B) (f (g y)) y) ->
- \(t : (x : A) -> IdP (<!0> A) (g (f x)) x) ->
- \(sA : (a b : A) -> (a0 b0 : IdP (<!0> A) a b) -> IdP (<!0> IdP (<!0> A) a b) a0 b0) ->
+ \(s : (y : B) -> PathP (<!0> B) (f (g y)) y) ->
+ \(t : (x : A) -> PathP (<!0> A) (g (f x)) x) ->
+ \(sA : (a b : A) -> (a0 b0 : PathP (<!0> A) a b) -> PathP (<!0> PathP (<!0> A) a b) a0 b0) ->
\(b0 b1 : B) ->
- \(p q : IdP (<!0> B) b0 b1) ->
+ \(p q : PathP (<!0> B) b0 b1) ->
<!1 !2> comp B (comp B (f (sA (g b0) (g b1) (<!1> comp A (g (p @ !1)) [ (!1 = 0) -> <!2> comp A (g (comp B b0 [ (!2 = 1) -> <!3> comp B (f (comp A (g b0) [ (!3 = 1) -> <!4> t (g b0) @ -!4 ])) [ (!3 = 0) -> <!4> s b0 @ !4, (!3 = 1) -> <!4> s (f (g b0)) @ !4 ] ])) [ (!2 = 1) -> <!3> t (g b0) @ !3 ], (!1 = 1) -> <!2> comp A (g (comp B b1 [ (!2 = 1) -> <!3> comp B (f (comp A (g b1) [ (!3 = 1) -> <!4> t (g b1) @ -!4 ])) [ (!3 = 0) -> <!4> s b1 @ !4, (!3 = 1) -> <!4> s (f (g b1)) @ !4 ] ])) [ (!2 = 1) -> <!3> t (g b1) @ !3 ] ]) (<!1> comp A (g (q @ !1)) [ (!1 = 0) -> <!2> comp A (g (comp B b0 [ (!2 = 1) -> <!3> comp B (f (comp A (g b0) [ (!3 = 1) -> <!4> t (g b0) @ -!4 ])) [ (!3 = 0) -> <!4> s b0 @ !4, (!3 = 1) -> <!4> s (f (g b0)) @ !4 ] ])) [ (!2 = 1) -> <!3> t (g b0) @ !3 ], (!1 = 1) -> <!2> comp A (g (comp B b1 [ (!2 = 1) -> <!3> comp B (f (comp A (g b1) [ (!3 = 1) -> <!4> t (g b1) @ -!4 ])) [ (!3 = 0) -> <!4> s b1 @ !4, (!3 = 1) -> <!4> s (f (g b1)) @ !4 ] ])) [ (!2 = 1) -> <!3> t (g b1) @ !3 ] ]) @ !1 @ !2)) [ (!2 = 0) -> <!3> comp B b0 [ (!3 = 0) -> <!4> comp B (f (comp A (g b0) [ (!4 = 1) -> <!5> t (g b0) @ -!5 ])) [ (!4 = 0) -> <!5> s b0 @ !5, (!4 = 1) -> <!5> s (f (g b0)) @ !5 ] ], (!2 = 1) -> <!3> comp B b1 [ (!3 = 0) -> <!4> comp B (f (comp A (g b1) [ (!4 = 1) -> <!5> t (g b1) @ -!5 ])) [ (!4 = 0) -> <!5> s b1 @ !5, (!4 = 1) -> <!5> s (f (g b1)) @ !5 ] ] ]) [ (!1 = 0) -> <!3> comp B (comp B (comp B (comp B (p @ !2) [ (!3 = 0) -> <!4> comp B (f (comp A (g (p @ !2)) [ (!4 = 1) -> <!5> t (g (p @ !2)) @ -!5 ])) [ (!4 = 0) -> <!5> s (p @ !2) @ !5, (!4 = 1) -> <!5> s (f (g (p @ !2))) @ !5 ] ]) [ (!2 = 0) -> <!4> comp B (comp B b0 [ (!3 = 0)(!4 = 1) -> <!5> comp B (f (comp A (g b0) [ (!5 = 1) -> <!6> t (g b0) @ -!6 ])) [ (!5 = 0) -> <!6> s b0 @ !6, (!5 = 1) -> <!6> s (f (g b0)) @ !6 ] ]) [ (!3 = 0) -> <!5> comp B (f (comp A (comp A (g (comp B b0 [ (!4 = 1) -> <!5> comp B (f (comp A (g b0) [ (!5 = 1) -> <!6> t (g b0) @ -!6 ])) [ (!5 = 0) -> <!6> s b0 @ !6, (!5 = 1) -> <!6> s (f (g b0)) @ !6 ] ])) [ (!4 = 1) -> <!6> t (g b0) @ (!5 /\ !6) ]) [ (!4 = 1) -> <!6> t (g b0) @ (!5 /\ -!6), (!5 = 1) -> <!6> t (comp A (g (comp B b0 [ (!4 = 1) -> <!5> comp B (f (comp A (g b0) [ (!5 = 1) -> <!6> t (g b0) @ -!6 ])) [ (!5 = 0) -> <!6> s b0 @ !6, (!5 = 1) -> <!6> s (f (g b0)) @ !6 ] ])) [ (!4 = 1) -> <!5> t (g b0) @ !5 ]) @ -!6 ])) [ (!4 = 1) -> <!6> s (f (g b0)) @ !6, (!5 = 0) -> <!6> s (comp B b0 [ (!4 = 1) -> <!5> comp B (f (comp A (g b0) [ (!5 = 1) -> <!6> t (g b0) @ -!6 ])) [ (!5 = 0) -> <!6> s b0 @ !6, (!5 = 1) -> <!6> s (f (g b0)) @ !6 ] ]) @ !6, (!5 = 1) -> <!6> s (f (comp A (g (comp B b0 [ (!4 = 1) -> <!5> comp B (f (comp A (g b0) [ (!5 = 1) -> <!6> t (g b0) @ -!6 ])) [ (!5 = 0) -> <!6> s b0 @ !6, (!5 = 1) -> <!6> s (f (g b0)) @ !6 ] ])) [ (!4 = 1) -> <!5> t (g b0) @ !5 ])) @ !6 ] ], (!2 = 1) -> <!4> comp B (comp B b1 [ (!3 = 0)(!4 = 1) -> <!5> comp B (f (comp A (g b1) [ (!5 = 1) -> <!6> t (g b1) @ -!6 ])) [ (!5 = 0) -> <!6> s b1 @ !6, (!5 = 1) -> <!6> s (f (g b1)) @ !6 ] ]) [ (!3 = 0) -> <!5> comp B (f (comp A (comp A (g (comp B b1 [ (!4 = 1) -> <!5> comp B (f (comp A (g b1) [ (!5 = 1) -> <!6> t (g b1) @ -!6 ])) [ (!5 = 0) -> <!6> s b1 @ !6, (!5 = 1) -> <!6> s (f (g b1)) @ !6 ] ])) [ (!4 = 1) -> <!6> t (g b1) @ (!5 /\ !6) ]) [ (!4 = 1) -> <!6> t (g b1) @ (!5 /\ -!6), (!5 = 1) -> <!6> t (comp A (g (comp B b1 [ (!4 = 1) -> <!5> comp B (f (comp A (g b1) [ (!5 = 1) -> <!6> t (g b1) @ -!6 ])) [ (!5 = 0) -> <!6> s b1 @ !6, (!5 = 1) -> <!6> s (f (g b1)) @ !6 ] ])) [ (!4 = 1) -> <!5> t (g b1) @ !5 ]) @ -!6 ])) [ (!4 = 1) -> <!6> s (f (g b1)) @ !6, (!5 = 0) -> <!6> s (comp B b1 [ (!4 = 1) -> <!5> comp B (f (comp A (g b1) [ (!5 = 1) -> <!6> t (g b1) @ -!6 ])) [ (!5 = 0) -> <!6> s b1 @ !6, (!5 = 1) -> <!6> s (f (g b1)) @ !6 ] ]) @ !6, (!5 = 1) -> <!6> s (f (comp A (g (comp B b1 [ (!4 = 1) -> <!5> comp B (f (comp A (g b1) [ (!5 = 1) -> <!6> t (g b1) @ -!6 ])) [ (!5 = 0) -> <!6> s b1 @ !6, (!5 = 1) -> <!6> s (f (g b1)) @ !6 ] ])) [ (!4 = 1) -> <!5> t (g b1) @ !5 ])) @ !6 ] ] ]) [ (!3 = 0) -> <!4> comp B (f (g (p @ !2))) [ (!2 = 0) -> <!5> f (comp A (g (comp B b0 [ (!5 = 1) -> <!6> comp B (f (comp A (g b0) [ (!6 = 1) -> <!7> t (g b0) @ -!7 ])) [ (!6 = 0) -> <!7> s b0 @ !7, (!6 = 1) -> <!7> s (f (g b0)) @ !7 ] ])) [ (!5 = 1) -> <!6> t (g b0) @ !6 ]), (!2 = 1) -> <!5> f (comp A (g (comp B b1 [ (!5 = 1) -> <!6> comp B (f (comp A (g b1) [ (!6 = 1) -> <!7> t (g b1) @ -!7 ])) [ (!6 = 0) -> <!7> s b1 @ !7, (!6 = 1) -> <!7> s (f (g b1)) @ !7 ] ])) [ (!5 = 1) -> <!6> t (g b1) @ !6 ]), (!4 = 0) -> <!5> comp B (f (g (p @ !2))) [ (!2 = 0) -> <!6> f (comp A (g (comp B b0 [ (!5 = 1)(!6 = 1) -> <!7> comp B (f (comp A (g b0) [ (!7 = 1) -> <!8> t (g b0) @ -!8 ])) [ (!7 = 0) -> <!8> s b0 @ !8, (!7 = 1) -> <!8> s (f (g b0)) @ !8 ] ])) [ (!5 = 1)(!6 = 1) -> <!7> t (g b0) @ !7 ]), (!2 = 1) -> <!6> f (comp A (g (comp B b1 [ (!5 = 1)(!6 = 1) -> <!7> comp B (f (comp A (g b1) [ (!7 = 1) -> <!8> t (g b1) @ -!8 ])) [ (!7 = 0) -> <!8> s b1 @ !8, (!7 = 1) -> <!8> s (f (g b1)) @ !8 ] ])) [ (!5 = 1)(!6 = 1) -> <!7> t (g b1) @ !7 ]) ], (!4 = 1) -> <!5> f (comp A (g (p @ !2)) [ (!2 = 0) -> <!6> comp A (g (comp B b0 [ (!5 = 1)(!6 = 1) -> <!7> comp B (f (comp A (g b0) [ (!7 = 1) -> <!8> t (g b0) @ -!8 ])) [ (!7 = 0) -> <!8> s b0 @ !8, (!7 = 1) -> <!8> s (f (g b0)) @ !8 ] ])) [ (!5 = 1)(!6 = 1) -> <!7> t (g b0) @ !7 ], (!2 = 1) -> <!6> comp A (g (comp B b1 [ (!5 = 1)(!6 = 1) -> <!7> comp B (f (comp A (g b1) [ (!7 = 1) -> <!8> t (g b1) @ -!8 ])) [ (!7 = 0) -> <!8> s b1 @ !8, (!7 = 1) -> <!8> s (f (g b1)) @ !8 ] ])) [ (!5 = 1)(!6 = 1) -> <!7> t (g b1) @ !7 ] ]) ] ]) [ (!2 = 0) -> <!4> comp B b0 [ (!3 = 0)(!4 = 0) -> <!5> comp B (f (comp A (g b0) [ (!5 = 1) -> <!6> t (g b0) @ -!6 ])) [ (!5 = 0) -> <!6> s b0 @ !6, (!5 = 1) -> <!6> s (f (g b0)) @ !6 ] ], (!2 = 1) -> <!4> comp B b1 [ (!3 = 0)(!4 = 0) -> <!5> comp B (f (comp A (g b1) [ (!5 = 1) -> <!6> t (g b1) @ -!6 ])) [ (!5 = 0) -> <!6> s b1 @ !6, (!5 = 1) -> <!6> s (f (g b1)) @ !6 ] ] ], (!1 = 1) -> <!3> comp B (comp B (comp B (comp B (q @ !2) [ (!3 = 0) -> <!4> comp B (f (comp A (g (q @ !2)) [ (!4 = 1) -> <!5> t (g (q @ !2)) @ -!5 ])) [ (!4 = 0) -> <!5> s (q @ !2) @ !5, (!4 = 1) -> <!5> s (f (g (q @ !2))) @ !5 ] ]) [ (!2 = 0) -> <!4> comp B (comp B b0 [ (!3 = 0)(!4 = 1) -> <!5> comp B (f (comp A (g b0) [ (!5 = 1) -> <!6> t (g b0) @ -!6 ])) [ (!5 = 0) -> <!6> s b0 @ !6, (!5 = 1) -> <!6> s (f (g b0)) @ !6 ] ]) [ (!3 = 0) -> <!5> comp B (f (comp A (comp A (g (comp B b0 [ (!4 = 1) -> <!5> comp B (f (comp A (g b0) [ (!5 = 1) -> <!6> t (g b0) @ -!6 ])) [ (!5 = 0) -> <!6> s b0 @ !6, (!5 = 1) -> <!6> s (f (g b0)) @ !6 ] ])) [ (!4 = 1) -> <!6> t (g b0) @ (!5 /\ !6) ]) [ (!4 = 1) -> <!6> t (g b0) @ (!5 /\ -!6), (!5 = 1) -> <!6> t (comp A (g (comp B b0 [ (!4 = 1) -> <!5> comp B (f (comp A (g b0) [ (!5 = 1) -> <!6> t (g b0) @ -!6 ])) [ (!5 = 0) -> <!6> s b0 @ !6, (!5 = 1) -> <!6> s (f (g b0)) @ !6 ] ])) [ (!4 = 1) -> <!5> t (g b0) @ !5 ]) @ -!6 ])) [ (!4 = 1) -> <!6> s (f (g b0)) @ !6, (!5 = 0) -> <!6> s (comp B b0 [ (!4 = 1) -> <!5> comp B (f (comp A (g b0) [ (!5 = 1) -> <!6> t (g b0) @ -!6 ])) [ (!5 = 0) -> <!6> s b0 @ !6, (!5 = 1) -> <!6> s (f (g b0)) @ !6 ] ]) @ !6, (!5 = 1) -> <!6> s (f (comp A (g (comp B b0 [ (!4 = 1) -> <!5> comp B (f (comp A (g b0) [ (!5 = 1) -> <!6> t (g b0) @ -!6 ])) [ (!5 = 0) -> <!6> s b0 @ !6, (!5 = 1) -> <!6> s (f (g b0)) @ !6 ] ])) [ (!4 = 1) -> <!5> t (g b0) @ !5 ])) @ !6 ] ], (!2 = 1) -> <!4> comp B (comp B b1 [ (!3 = 0)(!4 = 1) -> <!5> comp B (f (comp A (g b1) [ (!5 = 1) -> <!6> t (g b1) @ -!6 ])) [ (!5 = 0) -> <!6> s b1 @ !6, (!5 = 1) -> <!6> s (f (g b1)) @ !6 ] ]) [ (!3 = 0) -> <!5> comp B (f (comp A (comp A (g (comp B b1 [ (!4 = 1) -> <!5> comp B (f (comp A (g b1) [ (!5 = 1) -> <!6> t (g b1) @ -!6 ])) [ (!5 = 0) -> <!6> s b1 @ !6, (!5 = 1) -> <!6> s (f (g b1)) @ !6 ] ])) [ (!4 = 1) -> <!6> t (g b1) @ (!5 /\ !6) ]) [ (!4 = 1) -> <!6> t (g b1) @ (!5 /\ -!6), (!5 = 1) -> <!6> t (comp A (g (comp B b1 [ (!4 = 1) -> <!5> comp B (f (comp A (g b1) [ (!5 = 1) -> <!6> t (g b1) @ -!6 ])) [ (!5 = 0) -> <!6> s b1 @ !6, (!5 = 1) -> <!6> s (f (g b1)) @ !6 ] ])) [ (!4 = 1) -> <!5> t (g b1) @ !5 ]) @ -!6 ])) [ (!4 = 1) -> <!6> s (f (g b1)) @ !6, (!5 = 0) -> <!6> s (comp B b1 [ (!4 = 1) -> <!5> comp B (f (comp A (g b1) [ (!5 = 1) -> <!6> t (g b1) @ -!6 ])) [ (!5 = 0) -> <!6> s b1 @ !6, (!5 = 1) -> <!6> s (f (g b1)) @ !6 ] ]) @ !6, (!5 = 1) -> <!6> s (f (comp A (g (comp B b1 [ (!4 = 1) -> <!5> comp B (f (comp A (g b1) [ (!5 = 1) -> <!6> t (g b1) @ -!6 ])) [ (!5 = 0) -> <!6> s b1 @ !6, (!5 = 1) -> <!6> s (f (g b1)) @ !6 ] ])) [ (!4 = 1) -> <!5> t (g b1) @ !5 ])) @ !6 ] ] ]) [ (!3 = 0) -> <!4> comp B (f (g (q @ !2))) [ (!2 = 0) -> <!5> f (comp A (g (comp B b0 [ (!5 = 1) -> <!6> comp B (f (comp A (g b0) [ (!6 = 1) -> <!7> t (g b0) @ -!7 ])) [ (!6 = 0) -> <!7> s b0 @ !7, (!6 = 1) -> <!7> s (f (g b0)) @ !7 ] ])) [ (!5 = 1) -> <!6> t (g b0) @ !6 ]), (!2 = 1) -> <!5> f (comp A (g (comp B b1 [ (!5 = 1) -> <!6> comp B (f (comp A (g b1) [ (!6 = 1) -> <!7> t (g b1) @ -!7 ])) [ (!6 = 0) -> <!7> s b1 @ !7, (!6 = 1) -> <!7> s (f (g b1)) @ !7 ] ])) [ (!5 = 1) -> <!6> t (g b1) @ !6 ]), (!4 = 0) -> <!5> comp B (f (g (q @ !2))) [ (!2 = 0) -> <!6> f (comp A (g (comp B b0 [ (!5 = 1)(!6 = 1) -> <!7> comp B (f (comp A (g b0) [ (!7 = 1) -> <!8> t (g b0) @ -!8 ])) [ (!7 = 0) -> <!8> s b0 @ !8, (!7 = 1) -> <!8> s (f (g b0)) @ !8 ] ])) [ (!5 = 1)(!6 = 1) -> <!7> t (g b0) @ !7 ]), (!2 = 1) -> <!6> f (comp A (g (comp B b1 [ (!5 = 1)(!6 = 1) -> <!7> comp B (f (comp A (g b1) [ (!7 = 1) -> <!8> t (g b1) @ -!8 ])) [ (!7 = 0) -> <!8> s b1 @ !8, (!7 = 1) -> <!8> s (f (g b1)) @ !8 ] ])) [ (!5 = 1)(!6 = 1) -> <!7> t (g b1) @ !7 ]) ], (!4 = 1) -> <!5> f (comp A (g (q @ !2)) [ (!2 = 0) -> <!6> comp A (g (comp B b0 [ (!5 = 1)(!6 = 1) -> <!7> comp B (f (comp A (g b0) [ (!7 = 1) -> <!8> t (g b0) @ -!8 ])) [ (!7 = 0) -> <!8> s b0 @ !8, (!7 = 1) -> <!8> s (f (g b0)) @ !8 ] ])) [ (!5 = 1)(!6 = 1) -> <!7> t (g b0) @ !7 ], (!2 = 1) -> <!6> comp A (g (comp B b1 [ (!5 = 1)(!6 = 1) -> <!7> comp B (f (comp A (g b1) [ (!7 = 1) -> <!8> t (g b1) @ -!8 ])) [ (!7 = 0) -> <!8> s b1 @ !8, (!7 = 1) -> <!8> s (f (g b1)) @ !8 ] ])) [ (!5 = 1)(!6 = 1) -> <!7> t (g b1) @ !7 ] ]) ] ]) [ (!2 = 0) -> <!4> comp B b0 [ (!3 = 0)(!4 = 0) -> <!5> comp B (f (comp A (g b0) [ (!5 = 1) -> <!6> t (g b0) @ -!6 ])) [ (!5 = 0) -> <!6> s b0 @ !6, (!5 = 1) -> <!6> s (f (g b0)) @ !6 ] ], (!2 = 1) -> <!4> comp B b1 [ (!3 = 0)(!4 = 0) -> <!5> comp B (f (comp A (g b1) [ (!5 = 1) -> <!6> t (g b1) @ -!6 ])) [ (!5 = 0) -> <!6> s b1 @ !6, (!5 = 1) -> <!6> s (f (g b1)) @ !6 ] ] ] ]
-}
<i> [ (i = 0) -> inj a, (i = 1) -> inj b ]
quoteq' (A : U) (R : A -> A -> U) (a b : A) (r : R a b)
- : Id (Quot A R) (inj a) (inj b) = <i> quoteq {Quot A R} a b r @ i
+ : Path (Quot A R) (inj a) (inj b) = <i> quoteq {Quot A R} a b r @ i
-- Test to define circle as a quotient of unit
base -> inj tt
loop @ i -> quoteq{s1quot} tt tt tt @ i
-rem3 : (a : Unit) -> Id s1quot (inj tt) (inj a) = split
+rem3 : (a : Unit) -> Path s1quot (inj tt) (inj a) = split
tt -> <i> inj tt
-test : Id U s1quot S1 =
- isoId s1quot S1 f1 f2 rem1 rem2
+test : Path U s1quot S1 =
+ isoPath s1quot S1 f1 f2 rem1 rem2
where
- rem1 : (y : S1) -> Id S1 (f1 (f2 y)) y = split
+ rem1 : (y : S1) -> Path S1 (f1 (f2 y)) y = split
base -> <i> base
loop @ i -> <j> loop1 @ i
- rem2 : (x : s1quot) -> Id s1quot (f2 (f1 x)) x = split
+ rem2 : (x : s1quot) -> Path s1quot (f2 (f1 x)) x = split
inj a -> rem3 a
quoteq a b r @ i -> ?
data setquot (A : U) (R : A -> A -> U) =
quot (a : A)
| identification (a b : A) (r : R a b) <i> [ (i = 0) -> quot a, (i = 1) -> quot b ]
- | setTruncation (a b : setquot A R) (p q : Id (setquot A R) a b) <i j>
+ | setTruncation (a b : setquot A R) (p q : Path (setquot A R) a b) <i j>
[ (i = 0) -> p @ j
, (i = 1) -> q @ j
, (j = 0) -> a
-}
identsetquot (A : U) (R : A -> A -> U) (a b : A) (r : R a b)
- : Id (setquot A R) (quot a) (quot b) = <i> identification {setquot A R} a b r @ i
+ : Path (setquot A R) (quot a) (quot b) = <i> identification {setquot A R} a b r @ i
setsetquot (A : U) (R : A -> A -> U) : set (setquot A R) =
- \(a b : setquot A R) (p q : Id (setquot A R) a b) ->
+ \(a b : setquot A R) (p q : Path (setquot A R) a b) ->
<i j> setTruncation {setquot A R} a b p q @ i @ j
\r
import prelude\r
\r
-section (A B : U) (f : A -> B) (g : B -> A) :U = (b : B) -> Id B (f (g b)) b\r
+section (A B : U) (f : A -> B) (g : B -> A) :U = (b : B) -> Path B (f (g b)) b\r
\r
retract (A B : U) (f : A -> B) (g : B -> A) : U = section B A g f\r
\r
lemRetract (A B : U) (f : A -> B) (g : B -> A) (rfg : retract A B f g) (x y:A) :\r
- Id A (g (f x)) (g (f y)) -> Id A x y\r
+ Path A (g (f x)) (g (f y)) -> Path A x y\r
= compUp A (g (f x)) x (g (f y)) y (rfg x) (rfg y)\r
\r
retractProp (A B : U) (f : A -> B) (g : B -> A) (rfg : retract A B f g) (pB :prop B) (x y:A) \r
- : Id A x y = lemRetract A B f g rfg x y (<i> g (pB (f x) (f y) @ i))\r
+ : Path A x y = lemRetract A B f g rfg x y (<i> g (pB (f x) (f y) @ i))\r
\r
retractInv (A B : U) (f : A -> B) (g : B -> A) (rfg : retract A B f g)\r
- (x y : A) (q: Id B (f x) (f y)) : Id A x y = \r
+ (x y : A) (q: Path B (f x) (f y)) : Path A x y = \r
compUp A (g (f x)) x (g (f y)) y (rfg x) (rfg y) (<i> (g (q @ i)))\r
\r
--- lemRSquare (A B : U) (f : A -> B) (g : B -> A) (rfg: retract A B f g)(x y:A) (p : Id A x y) :\r
+-- lemRSquare (A B : U) (f : A -> B) (g : B -> A) (rfg: retract A B f g)(x y:A) (p : Path A x y) :\r
-- Square A (g (f x)) (g (f y)) (<i> g (f (p @ i))) x y\r
-- (retractInv A B f g rfg x y (<i> f (p@ i))) (rfg x) (rfg y) =\r
-- <j i> comp A (g (f (p @ j))) [(j=0) -> <l> (rfg x) @ (i/\l), (j=1) -> <l> (rfg y) @ (i/\l)]\r
\r
--- retractId (A B : U)(f : A -> B) (g : B -> A) (rfg : retract A B f g) (x y :A) (p:Id A x y) :\r
--- Id (Id A x y) (retractInv A B f g rfg x y (<i> f (p@ i))) p =\r
+-- retractPath (A B : U)(f : A -> B) (g : B -> A) (rfg : retract A B f g) (x y :A) (p:Path A x y) :\r
+-- Path (Path A x y) (retractInv A B f g rfg x y (<i> f (p@ i))) p =\r
-- <i j> comp A (g (f (p @ j))) [(j=0) -> rfg x,(j=1) -> rfg y,\r
-- (i=0) -> (lemRSquare A B f g rfg x y p) @ j,(i=1) -> rfg (p @ j)]\r
\r
-- retractSet (A B : U) (f : A -> B) (g : B -> A) (rfg : retract A B f g)\r
--- (sB : set B) (x y : A) : prop (Id A x y) =\r
--- retractProp (Id A x y) (Id B (f x) (f y)) (mapOnPath A B f x y)\r
--- (retractInv A B f g rfg x y) (retractId A B f g rfg x y) (sB (f x) (f y))\r
+-- (sB : set B) (x y : A) : prop (Path A x y) =\r
+-- retractProp (Path A x y) (Path B (f x) (f y)) (mapOnPath A B f x y)\r
+-- (retractInv A B f g rfg x y) (retractPath A B f g rfg x y) (sB (f x) (f y))\r
\r
import univalence
subtypeEquality (A : U) (B : A -> U) (pB : (x : A) -> prop (B x))
- (s t : (x : A) * B x) : Id A s.1 t.1 -> Id (Sigma A B) s t =
- trans (Id A s.1 t.1) (Id (Sigma A B) s t) rem
+ (s t : (x : A) * B x) : Path A s.1 t.1 -> Path (Sigma A B) s t =
+ trans (Path A s.1 t.1) (Path (Sigma A B) s t) rem
where
- rem : Id U (Id A s.1 t.1) (Id (Sigma A B) s t) =
+ rem : Path U (Path A s.1 t.1) (Path (Sigma A B) s t) =
<i> lemSigProp A B pB s t @ -i
-- (* Propositions *)
rem (y : B) : isContr (fiber A B f y) = (s,t)
where
s : fiber A B f y = (g y,pB y (f (g y)))
- t (w : fiber A B f y) : Id ((x : A) * Id B y (f x)) s w =
- subtypeEquality A (\(x : A) -> Id B y (f x)) pb s w r1
+ t (w : fiber A B f y) : Path ((x : A) * Path B y (f x)) s w =
+ subtypeEquality A (\(x : A) -> Path B y (f x)) pb s w r1
where
- pb (x : A) : (a b : Id B y (f x)) -> Id (Id B y (f x)) a b = propSet B pB y (f x)
- r1 : Id A s.1 w.1 = pA s.1 w.1
+ pb (x : A) : (a b : Path B y (f x)) -> Path (Path B y (f x)) a b = propSet B pB y (f x)
+ r1 : Path A s.1 w.1 = pA s.1 w.1
equivhProp (P P' : hProp) (f : P.1 -> P'.1) (g : P'.1 -> P.1) : equiv P.1 P'.1 =
(f,isEquivprop P.1 P'.1 f g P.2 P'.2)
-- Proof of uahp using full univalence
-uahp' (P P' : hProp) (f : P.1 -> P'.1) (g : P'.1 -> P.1) : Id hProp P P' =
+uahp' (P P' : hProp) (f : P.1 -> P'.1) (g : P'.1 -> P.1) : Path hProp P P' =
subtypeEquality U prop propIsProp P P' rem
where
- rem : Id U P.1 P'.1 = transport (<i> corrUniv P.1 P'.1 @ -i) (equivhProp P P' f g)
+ rem : Path U P.1 P'.1 = transport (<i> corrUniv P.1 P'.1 @ -i) (equivhProp P P' f g)
-- Direct proof of uahp
-uahp (P P' : hProp) (f : P.1 -> P'.1) (g : P'.1 -> P.1) : Id hProp P P' =
+uahp (P P' : hProp) (f : P.1 -> P'.1) (g : P'.1 -> P.1) : Path hProp P P' =
subtypeEquality U prop propIsProp P P' rem
where
- rem : Id U P.1 P'.1 = isoId P.1 P'.1 f g s t
- where s (y : P'.1) : Id P'.1 (f (g y)) y = P'.2 (f (g y)) y
- t (x : P.1) : Id P.1 (g (f x)) x = P.2 (g (f x)) x
+ rem : Path U P.1 P'.1 = isoPath P.1 P'.1 f g s t
+ where s (y : P'.1) : Path P'.1 (f (g y)) y = P'.2 (f (g y)) y
+ t (x : P.1) : Path P.1 (g (f x)) x = P.2 (g (f x)) x
-- A short proof that hProp form a set using univalence: (this is not needed!)
-propequiv (X Y : U) (H : prop Y) (f g : equiv X Y) : Id (equiv X Y) f g =
+propequiv (X Y : U) (H : prop Y) (f g : equiv X Y) : Path (equiv X Y) f g =
equivLemma X Y f g (<i> \(x : X) -> H (f.1 x) (g.1 x) @ i)
-propidU (X Y : U) : Id U X Y -> prop Y -> prop X = substInv U prop X Y
+propidU (X Y : U) : Path U X Y -> prop Y -> prop X = substInv U prop X Y
-sethProp (P P' : hProp) : prop (Id hProp P P') =
- propidU (Id hProp P P') (equiv P.1 P'.1) rem (propequiv P.1 P'.1 P'.2)
+sethProp (P P' : hProp) : prop (Path hProp P P') =
+ propidU (Path hProp P P') (equiv P.1 P'.1) rem (propequiv P.1 P'.1 P'.2)
where
- rem1 : Id U (Id hProp P P') (Id U P.1 P'.1) = lemSigProp U prop propIsProp P P'
- rem2 : Id U (Id U P.1 P'.1) (equiv P.1 P'.1) = corrUniv P.1 P'.1
- rem : Id U (Id hProp P P') (equiv P.1 P'.1) =
- compId U (Id hProp P P') (Id U P.1 P'.1) (equiv P.1 P'.1) rem1 rem2
+ rem1 : Path U (Path hProp P P') (Path U P.1 P'.1) = lemSigProp U prop propIsProp P P'
+ rem2 : Path U (Path U P.1 P'.1) (equiv P.1 P'.1) = corrUniv P.1 P'.1
+ rem : Path U (Path hProp P P') (equiv P.1 P'.1) =
+ compPath U (Path hProp P P') (Path U P.1 P'.1) (equiv P.1 P'.1) rem1 rem2
-- (* Sets *)
-- This proof is quite cool, but it looks ugly...
p2 (f g : (x1 x2 : X) -> (R x1 x2).1 -> (A x1).1 -> (A x2).1) :
- Id ((x1 x2 : X) -> (R x1 x2).1 -> (A x1).1 -> (A x2).1) f g =
+ Path ((x1 x2 : X) -> (R x1 x2).1 -> (A x1).1 -> (A x2).1) f g =
<i> \(x1 x2 : X) (h1 : (R x1 x2).1) (h2 : (A x1).1) ->
(A x2).2 (f x1 x2 h1 h2) (g x1 x2 h1 h2) @ i
p3 (f g : (x1 x2 : X) -> (A x1).1 -> (A x2).1 -> (R x1 x2).1) :
- Id ((x1 x2 : X) -> (A x1).1 -> (A x2).1 -> (R x1 x2).1) f g =
+ Path ((x1 x2 : X) -> (A x1).1 -> (A x2).1 -> (R x1 x2).1) f g =
<i> \(x1 x2 : X) (h1 : (A x1).1) (h2 : (A x2).1) ->
(R x1 x2).2 (f x1 x2 h1 h2) (g x1 x2 h1 h2) @ i
p3 (x1 x2 : X) (X1 : (A x1).1) (X2 : (A x2).1) : (R.1 x1 x2).1 = tax x1 X0 x2 (sax X0 x1 X1) X2
setquotl0 (X : U) (R : eqrel X) (c : setquot X R.1) (x : carrier X c.1) :
- Id (setquot X R.1) (setquotpr X R x.1) c = subtypeEquality (hsubtypes X) (iseqclass X R.1) p (setquotpr X R x.1) c rem
+ Path (setquot X R.1) (setquotpr X R x.1) c = subtypeEquality (hsubtypes X) (iseqclass X R.1) p (setquotpr X R x.1) c rem
where
p (A : hsubtypes X) : prop (iseqclass X R.1 A) = propiseqclass X R.1 A
- rem : Id (hsubtypes X) (setquotpr X R x.1).1 c.1 = <i> \(x : X) -> (rem' x) @ i -- inlined use of funext
- where rem' (a : X) : Id hProp ((setquotpr X R x.1).1 a) (c.1 a) =
+ rem : Path (hsubtypes X) (setquotpr X R x.1).1 c.1 = <i> \(x : X) -> (rem' x) @ i -- inlined use of funext
+ where rem' (a : X) : Path hProp ((setquotpr X R x.1).1 a) (c.1 a) =
uahp' ((setquotpr X R x.1).1 a) (c.1 a) l2r r2l -- This is where uahp appears
where
l2r (r : ((setquotpr X R x.1).1 a).1) : (c.1 a).1 = eqax1 X R.1 c.1 c.2 x.1 a r x.2
(ps : (x : X) -> (P (setquotpr X R x)).1) (c : setquot X R.1) : (P c).1 = hinhuniv (carrier X c.1) (P c) f rem
where
f (x : carrier X c.1) : (P c).1 =
- let e : Id (setquot X R.1) (setquotpr X R x.1) c = setquotl0 X R c x
+ let e : Path (setquot X R.1) (setquotpr X R x.1) c = setquotl0 X R c x
in subst (setquot X R.1) (\(w : setquot X R.1) -> (P w).1) (setquotpr X R x.1) c e (ps x.1)
rem : (ishinh (carrier X c.1)).1 = eqax0 X R.1 c.1 c.2
sB (x : hsubtypes X) : set (iseqclass X R x) = propSet (iseqclass X R x) (propiseqclass X R x)
iscompsetquotpr (X : U) (R : eqrel X) (x x' : X) (a : (R.1 x x').1) :
- Id (setquot X R.1) (setquotpr X R x) (setquotpr X R x') =
+ Path (setquot X R.1) (setquotpr X R x) (setquotpr X R x') =
subtypeEquality (hsubtypes X) (iseqclass X R.1) rem1 (setquotpr X R x) (setquotpr X R x') rem2
where
rem1 (x : hsubtypes X) : prop (iseqclass X R.1 x) = propiseqclass X R.1 x
- rem2 : Id (hsubtypes X) (setquotpr X R x).1 (setquotpr X R x').1 =
+ rem2 : Path (hsubtypes X) (setquotpr X R x).1 (setquotpr X R x').1 =
<i> \(x0 : X) -> rem x0 @ i
where
- rem (x0 : X) : Id hProp (R.1 x x0) (R.1 x' x0) = uahp' (R.1 x x0) (R.1 x' x0) f g
+ rem (x0 : X) : Path hProp (R.1 x x0) (R.1 x' x0) = uahp' (R.1 x x0) (R.1 x' x0) f g
where
f (r0 : (R.1 x x0).1) : (R.1 x' x0).1 =
eqreltrans X R x' x x0 (eqrelsymm X R x x' a) r0
eqreltrans X R x x' x0 a r0
weqpathsinsetquot (X : U) (R : eqrel X) (x x' : X) :
- equiv (R.1 x x').1 (Id (setquot X R.1) (setquotpr X R x) (setquotpr X R x')) =
+ equiv (R.1 x x').1 (Path (setquot X R.1) (setquotpr X R x) (setquotpr X R x')) =
(iscompsetquotpr X R x x',rem)
where
rem : isEquiv (R.1 x x').1
- (Id (setquot X R.1) (setquotpr X R x) (setquotpr X R x'))
+ (Path (setquot X R.1) (setquotpr X R x) (setquotpr X R x'))
(iscompsetquotpr X R x x') =
isEquivprop (R.1 x x').1
- (Id (setquot X R.1) (setquotpr X R x) (setquotpr X R x'))
+ (Path (setquot X R.1) (setquotpr X R x) (setquotpr X R x'))
(iscompsetquotpr X R x x')
g pA pB
- where g (e : Id (setquot X R.1) (setquotpr X R x) (setquotpr X R x')) :
+ where g (e : Path (setquot X R.1) (setquotpr X R x) (setquotpr X R x')) :
(R.1 x x').1 = transport (<i> (rem1 @ i).1) rem
where
rem : (R.1 x' x').1 = eqrelrefl X R x'
- rem2 : Id hProp (R.1 x x') (R.1 x' x') = <i> (e @ i).1 x'
- rem1 : Id hProp (R.1 x' x') (R.1 x x') = <i> rem2 @ -i
+ rem2 : Path hProp (R.1 x x') (R.1 x' x') = <i> (e @ i).1 x'
+ rem1 : Path hProp (R.1 x' x') (R.1 x x') = <i> rem2 @ -i
pA : prop (R.1 x x').1 = (R.1 x x').2
- pB : prop (Id (setquot X R.1) (setquotpr X R x) (setquotpr X R x')) =
+ pB : prop (Path (setquot X R.1) (setquotpr X R x) (setquotpr X R x')) =
setsetquot X R.1 (setquotpr X R x) (setquotpr X R x')
isdecprop (X : U) : U = and (prop X) (dec X)
rem1 : prop (prop X) = propIsProp X
rem2 : prop X -> prop (dec X) = propDec X
-isdeceqif (X : U) (h : (x x' : X) -> isdecprop (Id X x x')) : discrete X =
+isdeceqif (X : U) (h : (x x' : X) -> isdecprop (Path X x x')) : discrete X =
\(x x' : X) -> (h x x').2
propEquiv (X Y : U) (w : equiv X Y) : prop X -> prop Y = subst U prop X Y rem
where
- rem : Id U X Y = transport (<i> corrUniv X Y @ -i) w
+ rem : Path U X Y = transport (<i> corrUniv X Y @ -i) w
isdecpropweqf (X Y : U) (w : equiv X Y) (hX : isdecprop X) : isdecprop Y = (rem1,rem2 hX.2)
where
isdiscretesetquot (X : U) (R : eqrel X) (is : (x x' : X) -> isdecprop (R.1 x x').1) :
discrete (setquot X R.1) = isdeceqif (setquot X R.1) rem
where
- rem : (x x' : setquot X R.1) -> isdecprop (Id (setquot X R.1) x x') =
+ rem : (x x' : setquot X R.1) -> isdecprop (Path (setquot X R.1) x x') =
setquotuniv2prop X R
- (\(x0 x0' : setquot X R.1) -> (isdecprop (Id (setquot X R.1) x0 x0'),
- propisdecprop (Id (setquot X R.1) x0 x0'))) rem'
+ (\(x0 x0' : setquot X R.1) -> (isdecprop (Path (setquot X R.1) x0 x0'),
+ propisdecprop (Path (setquot X R.1) x0 x0'))) rem'
where
- rem' (x0 x0' : X) : isdecprop (Id (setquot X R.1) (setquotpr X R x0) (setquotpr X R x0')) =
- isdecpropweqf (R.1 x0 x0').1 (Id (setquot X R.1) (setquotpr X R x0) (setquotpr X R x0'))
+ rem' (x0 x0' : X) : isdecprop (Path (setquot X R.1) (setquotpr X R x0) (setquotpr X R x0')) =
+ isdecpropweqf (R.1 x0 x0').1 (Path (setquot X R.1) (setquotpr X R x0) (setquotpr X R x0'))
(weqpathsinsetquot X R x0 x0') (is x0 x0')
discretetobool (X : U) (h : discrete X) (x y : X) : bool = rem (h x y)
where
- rem : dec (Id X x y) -> bool = split
+ rem : dec (Path X x y) -> bool = split
inl _ -> true
inr _ -> false
R : eqrel bool = (r1,r2)
where
- r1 : hrel bool = \(x y : bool) -> (Id bool x y,setbool x y)
- r2 : iseqrel bool r1 = ((compId bool,refl bool),inv bool)
+ r1 : hrel bool = \(x y : bool) -> (Path bool x y,setbool x y)
+ r2 : iseqrel bool r1 = ((compPath bool,refl bool),inv bool)
bool' : U = setquot bool R.1
true' : bool' = setquotpr bool R true
false' : bool' = setquotpr bool R false
P' (t : bool') : hProp =
- hdisj (Id bool' t true') (Id bool' t false')
+ hdisj (Path bool' t true') (Path bool' t false')
K' (t : bool') : (P' t).1 = setquotunivprop bool R P' ps t
where
ps : (x : bool) -> (P' (setquotpr bool R x)).1 = split
- false -> hdisj_in2 (Id bool' false' true')
- (Id bool' false' false') (<_> false')
- true -> hdisj_in1 (Id bool' true' true')
- (Id bool' true' false') (<_> true')
+ false -> hdisj_in2 (Path bool' false' true')
+ (Path bool' false' false') (<_> false')
+ true -> hdisj_in1 (Path bool' true' true')
+ (Path bool' true' false') (<_> true')
-test : (P' true').1 = hdisj_in1 (Id bool' true' true')
- (Id bool' true' false') (<_> true')
+test : (P' true').1 = hdisj_in1 (Path bool' true' true')
+ (Path bool' true' false') (<_> true')
test' : (P' true').1 = K' true'
--- test'' : Id (P' true').1 test test' = (P' true').2 test test'
+-- test'' : Path (P' true').1 test test' = (P' true').2 test test'
-- These two terms are not convertible:
--- test' : Id (P' true').1 (K' true')
--- (hdisj_in1 (Id (setquot bool R.1) true' true') (Id (setquot bool R.1) true' false') (<_> true')) =
+-- test' : Path (P' true').1 (K' true')
+-- (hdisj_in1 (Path (setquot bool R.1) true' true') (Path (setquot bool R.1) true' false') (<_> true')) =
-- <_> K' true'
-- Another test:
-true'neqfalse' (H : Id bool' true' false') : N0 = falseNeqTrue rem1
+true'neqfalse' (H : Path bool' true' false') : N0 = falseNeqTrue rem1
where
- rem : Id U (Id bool true true) (Id bool false true) = <i> ((H @ i).1 true).1
- rem1 : Id bool false true = comp rem (<_> true) []
+ rem : Path U (Path bool true true) (Path bool false true) = <i> ((H @ i).1 true).1
+ rem1 : Path bool false true = comp rem (<_> true) []
-test1 (x : bool') (H1 : Id bool' x true') (H2 : Id bool' x false') : N0 = true'neqfalse' rem
+test1 (x : bool') (H1 : Path bool' x true') (H2 : Path bool' x false') : N0 = true'neqfalse' rem
where
- rem : Id bool' true' false' = <i> comp (<_> bool') x [(i = 0) -> H1, (i = 1) -> H2]
+ rem : Path bool' true' false' = <i> comp (<_> bool') x [(i = 0) -> H1, (i = 1) -> H2]
-test2 (x : bool') (p1 : (ishinh (Id bool' x true')).1)
- (p2 : (ishinh (Id bool' x false')).1) : N0 =
- hinhuniv (Id bool' x true') (N0,propN0) rem p1
+test2 (x : bool') (p1 : (ishinh (Path bool' x true')).1)
+ (p2 : (ishinh (Path bool' x false')).1) : N0 =
+ hinhuniv (Path bool' x true') (N0,propN0) rem p1
where
- rem (H1 : Id bool' x true') : N0 =
- hinhuniv (Id bool' x false') (N0,propN0)
- (\(H2 : Id bool' x false') -> test1 x H1 H2) p2
+ rem (H1 : Path bool' x true') : N0 =
+ hinhuniv (Path bool' x false') (N0,propN0)
+ (\(H2 : Path bool' x false') -> test1 x H1 H2) p2
-- shorthand for this big type
-T (x : bool') : U = or (ishinh (Id bool' x true')).1 (ishinh (Id bool' x false')).1
+T (x : bool') : U = or (ishinh (Path bool' x true')).1 (ishinh (Path bool' x false')).1
-- test3 (x : bool') : prop (T x)
-test3 (x : bool') : (a b : T x) -> Id (T x) a b = split
+test3 (x : bool') : (a b : T x) -> Path (T x) a b = split
inl a' -> rem
where
- rem : (b : T x) -> Id (T x) (inl a') b = split
- inl b' -> <i> inl (propishinh (Id bool' x true') a' b' @ i)
- inr b' -> efq (Id (T x) (inl a') (inr b')) (test2 x a' b')
+ rem : (b : T x) -> Path (T x) (inl a') b = split
+ inl b' -> <i> inl (propishinh (Path bool' x true') a' b' @ i)
+ inr b' -> efq (Path (T x) (inl a') (inr b')) (test2 x a' b')
inr a' -> rem
where
- rem : (b : T x) -> Id (T x) (inr a') b = split
- inl b' -> efq (Id (T x) (inr a') (inl b')) (test2 x b' a')
- inr b' -> <i> inr (propishinh (Id bool' x false') a' b' @ i)
+ rem : (b : T x) -> Path (T x) (inr a') b = split
+ inl b' -> efq (Path (T x) (inr a') (inl b')) (test2 x b' a')
+ inr b' -> <i> inr (propishinh (Path bool' x false') a' b' @ i)
-f (x : bool') : or (ishinh (Id bool' x true')).1 (ishinh (Id bool' x false')).1 -> bool = split
+f (x : bool') : or (ishinh (Path bool' x true')).1 (ishinh (Path bool' x false')).1 -> bool = split
inl _ -> true
inr _ -> false
-bar (x : bool') : or (Id bool' x true') (Id bool' x false') ->
- or (ishinh (Id bool' x true')).1 (ishinh (Id bool' x false')).1 = split
- inl p -> inl (hinhpr (Id bool' x true') p)
- inr p -> inr (hinhpr (Id bool' x false') p)
+bar (x : bool') : or (Path bool' x true') (Path bool' x false') ->
+ or (ishinh (Path bool' x true')).1 (ishinh (Path bool' x false')).1 = split
+ inl p -> inl (hinhpr (Path bool' x true') p)
+ inr p -> inr (hinhpr (Path bool' x false') p)
-- finally the map:
foo (x : bool') (x' : (P' x).1) : bool = f x rem
where
- rem : or (ishinh (Id bool' x true')).1 (ishinh (Id bool' x false')).1 =
- hinhuniv (or (Id bool' x true') (Id bool' x false'))
- (or (ishinh (Id bool' x true')).1 (ishinh (Id bool' x false')).1,test3 x)
+ rem : or (ishinh (Path bool' x true')).1 (ishinh (Path bool' x false')).1 =
+ hinhuniv (or (Path bool' x true') (Path bool' x false'))
+ (or (ishinh (Path bool' x true')).1 (ishinh (Path bool' x false')).1,test3 x)
(bar x) x'
-- > :n testfoo
-- Time: 0m0.490s
testfoo : bool = foo true' (K' true')
-testfoo' : Id bool (foo true' (K' true')) true = <i> foo true' (K' true')
+testfoo' : Path bool (foo true' (K' true')) true = <i> foo true' (K' true')
-- Tests of checking normal forms:
-ntrue' : bool' = (\(x : bool) -> (IdP (<i0> bool) true x,lem8 x),((\(P : Sigma U (\(X : U) -> (a b : X) -> IdP (<i0> X) a b)) -> \(f : (Sigma bool (\(x : bool) -> IdP (<i0> bool) true x)) -> P.1) -> f ((true,<i0> true)),\(x1 x2 : bool) -> \(X1 : IdP (<i0> bool) x1 x2) -> \(X2 : IdP (<i0> bool) true x1) -> <i0> comp (<i1> bool) (X2 @ i0) [ (i0 = 0) -> <i1> true, (i0 = 1) -> <i1> X1 @ i1 ]),\(x1 x2 : bool) -> \(X1 : IdP (<i0> bool) true x1) -> \(X2 : IdP (<i0> bool) true x2) -> <i0> comp (<i1> bool) (X1 @ -i0) [ (i0 = 0) -> <i1> x1, (i0 = 1) -> <i1> X2 @ i1 ]))
+ntrue' : bool' = (\(x : bool) -> (PathP (<i0> bool) true x,lem8 x),((\(P : Sigma U (\(X : U) -> (a b : X) -> PathP (<i0> X) a b)) -> \(f : (Sigma bool (\(x : bool) -> PathP (<i0> bool) true x)) -> P.1) -> f ((true,<i0> true)),\(x1 x2 : bool) -> \(X1 : PathP (<i0> bool) x1 x2) -> \(X2 : PathP (<i0> bool) true x1) -> <i0> comp (<i1> bool) (X2 @ i0) [ (i0 = 0) -> <i1> true, (i0 = 1) -> <i1> X1 @ i1 ]),\(x1 x2 : bool) -> \(X1 : PathP (<i0> bool) true x1) -> \(X2 : PathP (<i0> bool) true x2) -> <i0> comp (<i1> bool) (X1 @ -i0) [ (i0 = 0) -> <i1> x1, (i0 = 1) -> <i1> X2 @ i1 ]))
nhdisj_in1 : (P Q : U) (X : P) -> (hdisj P Q).1 =
- \(P Q : U) -> \(X : P) -> \(P0 : Sigma U (\(X0 : U) -> (a b : X0) -> IdP (<!0> X0) a b)) -> \(f : or P Q -> P0.1) -> f (inl X)
+ \(P Q : U) -> \(X : P) -> \(P0 : Sigma U (\(X0 : U) -> (a b : X0) -> PathP (<!0> X0) a b)) -> \(f : or P Q -> P0.1) -> f (inl X)
-ntest : (P' true').1 = \(P : Sigma U (\(X : U) -> (a b : X) -> IdP (<!0> X) a b)) -> \(f : or (IdP (<!0> Sigma (bool -> (Sigma U (\(X : U) -> (a b : X) -> IdP (<!0> X) a b))) (\(A : bool -> (Sigma U (\(X : U) -> (a b : X) -> IdP (<!0> X) a b))) -> Sigma (Sigma ((P0 : Sigma U (\(X : U) -> (a b : X) -> IdP (<!0> X) a b)) -> ((Sigma bool (\(x : bool) -> (A x).1)) -> P0.1) -> P0.1) (\(_ : (P0 : Sigma U (\(X : U) -> (a b : X) -> IdP (<!0> X) a b)) -> ((Sigma bool (\(x : bool) -> (A x).1)) -> P0.1) -> P0.1) -> (x1 x2 : bool) -> (IdP (<!0> bool) x1 x2) -> ((A x1).1 -> (A x2).1))) (\(_ : Sigma ((P0 : Sigma U (\(X : U) -> (a b : X) -> IdP (<!0> X) a b)) -> ((Sigma bool (\(x : bool) -> (A x).1)) -> P0.1) -> P0.1) (\(_ : (P0 : Sigma U (\(X : U) -> (a b : X) -> IdP (<!0> X) a b)) -> ((Sigma bool (\(x : bool) -> (A x).1)) -> P0.1) -> P0.1) -> (x1 x2 : bool) -> (IdP (<!0> bool) x1 x2) -> ((A x1).1 -> (A x2).1))) -> (x1 x2 : bool) -> (A x1).1 -> ((A x2).1 -> (IdP (<!0> bool) x1 x2))))) ((\(x : bool) -> (IdP (<!0> bool) true x,lem8 x),((\(P0 : Sigma U (\(X : U) -> (a b : X) -> IdP (<!0> X) a b)) -> \(f : (Sigma bool (\(x : bool) -> IdP (<!0> bool) true x)) -> P0.1) -> f ((true,<!0> true)),\(x1 x2 : bool) -> \(X1 : IdP (<!0> bool) x1 x2) -> \(X2 : IdP (<!0> bool) true x1) -> <!0> comp (<!1> bool) (X2 @ !0) [ (!0 = 0) -> <!1> true, (!0 = 1) -> <!1> X1 @ !1 ]),\(x1 x2 : bool) -> \(X1 : IdP (<!0> bool) true x1) -> \(X2 : IdP (<!0> bool) true x2) -> <!0> comp (<!1> bool) (X1 @ -!0) [ (!0 = 0) -> <!1> x1, (!0 = 1) -> <!1> X2 @ !1 ]))) ((\(x : bool) -> (IdP (<!0> bool) true x,lem8 x),((\(P0 : Sigma U (\(X : U) -> (a b : X) -> IdP (<!0> X) a b)) -> \(f : (Sigma bool (\(x : bool) -> IdP (<!0> bool) true x)) -> P0.1) -> f ((true,<!0> true)),\(x1 x2 : bool) -> \(X1 : IdP (<!0> bool) x1 x2) -> \(X2 : IdP (<!0> bool) true x1) -> <!0> comp (<!1> bool) (X2 @ !0) [ (!0 = 0) -> <!1> true, (!0 = 1) -> <!1> X1 @ !1 ]),\(x1 x2 : bool) -> \(X1 : IdP (<!0> bool) true x1) -> \(X2 : IdP (<!0> bool) true x2) -> <!0> comp (<!1> bool) (X1 @ -!0) [ (!0 = 0) -> <!1> x1, (!0 = 1) -> <!1> X2 @ !1 ])))) IdP (<!0> Sigma (bool -> (Sigma U (\(X : U) -> (a b : X) -> IdP (<!0> X) a b))) (\(A : bool -> (Sigma U (\(X : U) -> (a b : X) -> IdP (<!0> X) a b))) -> Sigma (Sigma ((P0 : Sigma U (\(X : U) -> (a b : X) -> IdP (<!0> X) a b)) -> ((Sigma bool (\(x : bool) -> (A x).1)) -> P0.1) -> P0.1) (\(_ : (P0 : Sigma U (\(X : U) -> (a b : X) -> IdP (<!0> X) a b)) -> ((Sigma bool (\(x : bool) -> (A x).1)) -> P0.1) -> P0.1) -> (x1 x2 : bool) -> (IdP (<!0> bool) x1 x2) -> ((A x1).1 -> (A x2).1))) (\(_ : Sigma ((P0 : Sigma U (\(X : U) -> (a b : X) -> IdP (<!0> X) a b)) -> ((Sigma bool (\(x : bool) -> (A x).1)) -> P0.1) -> P0.1) (\(_ : (P0 : Sigma U (\(X : U) -> (a b : X) -> IdP (<!0> X) a b)) -> ((Sigma bool (\(x : bool) -> (A x).1)) -> P0.1) -> P0.1) -> (x1 x2 : bool) -> (IdP (<!0> bool) x1 x2) -> ((A x1).1 -> (A x2).1))) -> (x1 x2 : bool) -> (A x1).1 -> ((A x2).1 -> (IdP (<!0> bool) x1 x2))))) ((\(x : bool) -> (IdP (<!0> bool) true x,lem8 x),((\(P0 : Sigma U (\(X : U) -> (a b : X) -> IdP (<!0> X) a b)) -> \(f : (Sigma bool (\(x : bool) -> IdP (<!0> bool) true x)) -> P0.1) -> f ((true,<!0> true)),\(x1 x2 : bool) -> \(X1 : IdP (<!0> bool) x1 x2) -> \(X2 : IdP (<!0> bool) true x1) -> <!0> comp (<!1> bool) (X2 @ !0) [ (!0 = 0) -> <!1> true, (!0 = 1) -> <!1> X1 @ !1 ]),\(x1 x2 : bool) -> \(X1 : IdP (<!0> bool) true x1) -> \(X2 : IdP (<!0> bool) true x2) -> <!0> comp (<!1> bool) (X1 @ -!0) [ (!0 = 0) -> <!1> x1, (!0 = 1) -> <!1> X2 @ !1 ]))) ((\(x : bool) -> (IdP (<!0> bool) false x,lem7 x),((\(P0 : Sigma U (\(X : U) -> (a b : X) -> IdP (<!0> X) a b)) -> \(f : (Sigma bool (\(x : bool) -> IdP (<!0> bool) false x)) -> P0.1) -> f ((false,<!0> false)),\(x1 x2 : bool) -> \(X1 : IdP (<!0> bool) x1 x2) -> \(X2 : IdP (<!0> bool) false x1) -> <!0> comp (<!1> bool) (X2 @ !0) [ (!0 = 0) -> <!1> false, (!0 = 1) -> <!1> X1 @ !1 ]),\(x1 x2 : bool) -> \(X1 : IdP (<!0> bool) false x1) -> \(X2 : IdP (<!0> bool) false x2) -> <!0> comp (<!1> bool) (X1 @ -!0) [ (!0 = 0) -> <!1> x1, (!0 = 1) -> <!1> X2 @ !1 ]))) -> P.1) -> f (inl (<!0> (\(x : bool) -> (IdP (<!0> bool) true x,lem8 x),((\(P0 : Sigma U (\(X : U) -> (a b : X) -> IdP (<!0> X) a b)) -> \(f0 : (Sigma bool (\(x : bool) -> IdP (<!0> bool) true x)) -> P0.1) -> f0 ((true,<!0> true)),\(x1 x2 : bool) -> \(X1 : IdP (<!0> bool) x1 x2) -> \(X2 : IdP (<!0> bool) true x1) -> <!0> comp (<!1> bool) (X2 @ !0) [ (!0 = 0) -> <!1> true, (!0 = 1) -> <!1> X1 @ !1 ]),\(x1 x2 : bool) -> \(X1 : IdP (<!0> bool) true x1) -> \(X2 : IdP (<!0> bool) true x2) -> <!0> comp (<!1> bool) (X1 @ -!0) [ (!0 = 0) -> <!1> x1, (!0 = 1) -> <!1> X2 @ !1 ]))))
+ntest : (P' true').1 = \(P : Sigma U (\(X : U) -> (a b : X) -> PathP (<!0> X) a b)) -> \(f : or (PathP (<!0> Sigma (bool -> (Sigma U (\(X : U) -> (a b : X) -> PathP (<!0> X) a b))) (\(A : bool -> (Sigma U (\(X : U) -> (a b : X) -> PathP (<!0> X) a b))) -> Sigma (Sigma ((P0 : Sigma U (\(X : U) -> (a b : X) -> PathP (<!0> X) a b)) -> ((Sigma bool (\(x : bool) -> (A x).1)) -> P0.1) -> P0.1) (\(_ : (P0 : Sigma U (\(X : U) -> (a b : X) -> PathP (<!0> X) a b)) -> ((Sigma bool (\(x : bool) -> (A x).1)) -> P0.1) -> P0.1) -> (x1 x2 : bool) -> (PathP (<!0> bool) x1 x2) -> ((A x1).1 -> (A x2).1))) (\(_ : Sigma ((P0 : Sigma U (\(X : U) -> (a b : X) -> PathP (<!0> X) a b)) -> ((Sigma bool (\(x : bool) -> (A x).1)) -> P0.1) -> P0.1) (\(_ : (P0 : Sigma U (\(X : U) -> (a b : X) -> PathP (<!0> X) a b)) -> ((Sigma bool (\(x : bool) -> (A x).1)) -> P0.1) -> P0.1) -> (x1 x2 : bool) -> (PathP (<!0> bool) x1 x2) -> ((A x1).1 -> (A x2).1))) -> (x1 x2 : bool) -> (A x1).1 -> ((A x2).1 -> (PathP (<!0> bool) x1 x2))))) ((\(x : bool) -> (PathP (<!0> bool) true x,lem8 x),((\(P0 : Sigma U (\(X : U) -> (a b : X) -> PathP (<!0> X) a b)) -> \(f : (Sigma bool (\(x : bool) -> PathP (<!0> bool) true x)) -> P0.1) -> f ((true,<!0> true)),\(x1 x2 : bool) -> \(X1 : PathP (<!0> bool) x1 x2) -> \(X2 : PathP (<!0> bool) true x1) -> <!0> comp (<!1> bool) (X2 @ !0) [ (!0 = 0) -> <!1> true, (!0 = 1) -> <!1> X1 @ !1 ]),\(x1 x2 : bool) -> \(X1 : PathP (<!0> bool) true x1) -> \(X2 : PathP (<!0> bool) true x2) -> <!0> comp (<!1> bool) (X1 @ -!0) [ (!0 = 0) -> <!1> x1, (!0 = 1) -> <!1> X2 @ !1 ]))) ((\(x : bool) -> (PathP (<!0> bool) true x,lem8 x),((\(P0 : Sigma U (\(X : U) -> (a b : X) -> PathP (<!0> X) a b)) -> \(f : (Sigma bool (\(x : bool) -> PathP (<!0> bool) true x)) -> P0.1) -> f ((true,<!0> true)),\(x1 x2 : bool) -> \(X1 : PathP (<!0> bool) x1 x2) -> \(X2 : PathP (<!0> bool) true x1) -> <!0> comp (<!1> bool) (X2 @ !0) [ (!0 = 0) -> <!1> true, (!0 = 1) -> <!1> X1 @ !1 ]),\(x1 x2 : bool) -> \(X1 : PathP (<!0> bool) true x1) -> \(X2 : PathP (<!0> bool) true x2) -> <!0> comp (<!1> bool) (X1 @ -!0) [ (!0 = 0) -> <!1> x1, (!0 = 1) -> <!1> X2 @ !1 ])))) PathP (<!0> Sigma (bool -> (Sigma U (\(X : U) -> (a b : X) -> PathP (<!0> X) a b))) (\(A : bool -> (Sigma U (\(X : U) -> (a b : X) -> PathP (<!0> X) a b))) -> Sigma (Sigma ((P0 : Sigma U (\(X : U) -> (a b : X) -> PathP (<!0> X) a b)) -> ((Sigma bool (\(x : bool) -> (A x).1)) -> P0.1) -> P0.1) (\(_ : (P0 : Sigma U (\(X : U) -> (a b : X) -> PathP (<!0> X) a b)) -> ((Sigma bool (\(x : bool) -> (A x).1)) -> P0.1) -> P0.1) -> (x1 x2 : bool) -> (PathP (<!0> bool) x1 x2) -> ((A x1).1 -> (A x2).1))) (\(_ : Sigma ((P0 : Sigma U (\(X : U) -> (a b : X) -> PathP (<!0> X) a b)) -> ((Sigma bool (\(x : bool) -> (A x).1)) -> P0.1) -> P0.1) (\(_ : (P0 : Sigma U (\(X : U) -> (a b : X) -> PathP (<!0> X) a b)) -> ((Sigma bool (\(x : bool) -> (A x).1)) -> P0.1) -> P0.1) -> (x1 x2 : bool) -> (PathP (<!0> bool) x1 x2) -> ((A x1).1 -> (A x2).1))) -> (x1 x2 : bool) -> (A x1).1 -> ((A x2).1 -> (PathP (<!0> bool) x1 x2))))) ((\(x : bool) -> (PathP (<!0> bool) true x,lem8 x),((\(P0 : Sigma U (\(X : U) -> (a b : X) -> PathP (<!0> X) a b)) -> \(f : (Sigma bool (\(x : bool) -> PathP (<!0> bool) true x)) -> P0.1) -> f ((true,<!0> true)),\(x1 x2 : bool) -> \(X1 : PathP (<!0> bool) x1 x2) -> \(X2 : PathP (<!0> bool) true x1) -> <!0> comp (<!1> bool) (X2 @ !0) [ (!0 = 0) -> <!1> true, (!0 = 1) -> <!1> X1 @ !1 ]),\(x1 x2 : bool) -> \(X1 : PathP (<!0> bool) true x1) -> \(X2 : PathP (<!0> bool) true x2) -> <!0> comp (<!1> bool) (X1 @ -!0) [ (!0 = 0) -> <!1> x1, (!0 = 1) -> <!1> X2 @ !1 ]))) ((\(x : bool) -> (PathP (<!0> bool) false x,lem7 x),((\(P0 : Sigma U (\(X : U) -> (a b : X) -> PathP (<!0> X) a b)) -> \(f : (Sigma bool (\(x : bool) -> PathP (<!0> bool) false x)) -> P0.1) -> f ((false,<!0> false)),\(x1 x2 : bool) -> \(X1 : PathP (<!0> bool) x1 x2) -> \(X2 : PathP (<!0> bool) false x1) -> <!0> comp (<!1> bool) (X2 @ !0) [ (!0 = 0) -> <!1> false, (!0 = 1) -> <!1> X1 @ !1 ]),\(x1 x2 : bool) -> \(X1 : PathP (<!0> bool) false x1) -> \(X2 : PathP (<!0> bool) false x2) -> <!0> comp (<!1> bool) (X1 @ -!0) [ (!0 = 0) -> <!1> x1, (!0 = 1) -> <!1> X2 @ !1 ]))) -> P.1) -> f (inl (<!0> (\(x : bool) -> (PathP (<!0> bool) true x,lem8 x),((\(P0 : Sigma U (\(X : U) -> (a b : X) -> PathP (<!0> X) a b)) -> \(f0 : (Sigma bool (\(x : bool) -> PathP (<!0> bool) true x)) -> P0.1) -> f0 ((true,<!0> true)),\(x1 x2 : bool) -> \(X1 : PathP (<!0> bool) x1 x2) -> \(X2 : PathP (<!0> bool) true x1) -> <!0> comp (<!1> bool) (X2 @ !0) [ (!0 = 0) -> <!1> true, (!0 = 1) -> <!1> X1 @ !1 ]),\(x1 x2 : bool) -> \(X1 : PathP (<!0> bool) true x1) -> \(X2 : PathP (<!0> bool) true x2) -> <!0> comp (<!1> bool) (X1 @ -!0) [ (!0 = 0) -> <!1> x1, (!0 = 1) -> <!1> X2 @ !1 ]))))
import equiv
-lemIdSig (A:U) (B : A -> U) (t u : Sigma A B) :
- Id U (Id (Sigma A B) t u) ((p : Id A t.1 u.1) * IdP (<i> B (p @ i)) t.2 u.2) =
- isoId T0 T1 f g s t where
- T0 : U = Id (Sigma A B) t u
- T1 : U = (p:Id A t.1 u.1) * IdP (<i> B (p@i)) t.2 u.2
+lemPathSig (A:U) (B : A -> U) (t u : Sigma A B) :
+ Path U (Path (Sigma A B) t u) ((p : Path A t.1 u.1) * PathP (<i> B (p @ i)) t.2 u.2) =
+ isoPath T0 T1 f g s t where
+ T0 : U = Path (Sigma A B) t u
+ T1 : U = (p:Path A t.1 u.1) * PathP (<i> B (p@i)) t.2 u.2
f (q:T0) : T1 = (<i> (q@i).1,<i> (q@i).2)
g (z:T1) : T0 = <i>(z.1 @i,z.2 @i)
- s (z:T1) : Id T1 (f (g z)) z = refl T1 z
- t (q:T0) : Id T0 (g (f q)) q = refl T0 q
+ s (z:T1) : Path T1 (f (g z)) z = refl T1 z
+ t (q:T0) : Path T0 (g (f q)) q = refl T0 q
-lemIdAnd (A B : U) (t u : and A B) :
- Id U (Id (and A B) t u) (and (Id A t.1 u.1) (Id B t.2 u.2)) = lemIdSig A (\(_ : A) -> B) t u
+lemPathAnd (A B : U) (t u : and A B) :
+ Path U (Path (and A B) t u) (and (Path A t.1 u.1) (Path B t.2 u.2)) = lemPathSig A (\(_ : A) -> B) t u
-lemTransp (A:U) (a:A) : Id A a (transport (<_>A) a) = fill (<_>A) a []
+lemTransp (A:U) (a:A) : Path A a (transport (<_>A) a) = fill (<_>A) a []
-funDepTr (A:U) (P:A->U) (a0 a1 :A) (p:Id A a0 a1) (u0:P a0) (u1:P a1) :
- Id U (IdP (<i> P (p@i)) u0 u1) (Id (P a1) (transport (<i> P (p@i)) u0) u1) =
- <j>IdP (<i>P (p@j\/i)) (comp (<i>P (p@j/\i)) u0 [(j=0)-><_>u0]) u1
+funDepTr (A:U) (P:A->U) (a0 a1 :A) (p:Path A a0 a1) (u0:P a0) (u1:P a1) :
+ Path U (PathP (<i> P (p@i)) u0 u1) (Path (P a1) (transport (<i> P (p@i)) u0) u1) =
+ <j>PathP (<i>P (p@j\/i)) (comp (<i>P (p@j/\i)) u0 [(j=0)-><_>u0]) u1
-lem2 (A:U) (B:A-> U) (t u : Sigma A B) (p:Id A t.1 u.1) :
- Id U (IdP (<i>B (p@i)) t.2 u.2) (Id (B u.1) (transport (<i>B (p@i)) t.2) u.2) =
+lem2 (A:U) (B:A-> U) (t u : Sigma A B) (p:Path A t.1 u.1) :
+ Path U (PathP (<i>B (p@i)) t.2 u.2) (Path (B u.1) (transport (<i>B (p@i)) t.2) u.2) =
funDepTr A B t.1 u.1 p t.2 u.2
-corSigProp (A:U) (B:A-> U) (pB : (x:A) -> prop (B x)) (t u : Sigma A B) (p:Id A t.1 u.1) :
- prop (IdP (<i>B (p@i)) t.2 u.2) = substInv U prop T0 T1 rem rem1
- where P : Id U (B t.1) (B u.1) = <i>B (p@i)
- T0 : U = IdP P t.2 u.2
- T1 : U = Id (B u.1) (transport P t.2) u.2
- rem : Id U T0 T1 = lem2 A B t u p -- funDepTr (B t.1) (B u.1) P t.2 u.2
+corSigProp (A:U) (B:A-> U) (pB : (x:A) -> prop (B x)) (t u : Sigma A B) (p:Path A t.1 u.1) :
+ prop (PathP (<i>B (p@i)) t.2 u.2) = substInv U prop T0 T1 rem rem1
+ where P : Path U (B t.1) (B u.1) = <i>B (p@i)
+ T0 : U = PathP P t.2 u.2
+ T1 : U = Path (B u.1) (transport P t.2) u.2
+ rem : Path U T0 T1 = lem2 A B t u p -- funDepTr (B t.1) (B u.1) P t.2 u.2
v2 : B u.1 = transport P t.2
rem1 : prop T1 = propSet (B u.1) (pB u.1) v2 u.2
-corSigSet (A:U) (B:A-> U) (sB : (x:A) -> set (B x)) (t u : Sigma A B) (p:Id A t.1 u.1) :
- prop (IdP (<i>B (p@i)) t.2 u.2) = substInv U prop T0 T1 rem rem1
- where P : Id U (B t.1) (B u.1) = <i>B (p@i)
- T0 : U = IdP P t.2 u.2
- T1 : U = Id (B u.1) (transport P t.2) u.2
- rem : Id U T0 T1 = lem2 A B t u p -- funDepTr (B t.1) (B u.1) P t.2 u.2
+corSigSet (A:U) (B:A-> U) (sB : (x:A) -> set (B x)) (t u : Sigma A B) (p:Path A t.1 u.1) :
+ prop (PathP (<i>B (p@i)) t.2 u.2) = substInv U prop T0 T1 rem rem1
+ where P : Path U (B t.1) (B u.1) = <i>B (p@i)
+ T0 : U = PathP P t.2 u.2
+ T1 : U = Path (B u.1) (transport P t.2) u.2
+ rem : Path U T0 T1 = lem2 A B t u p -- funDepTr (B t.1) (B u.1) P t.2 u.2
v2 : B u.1 = transport P t.2
rem1 : prop T1 = sB u.1 v2 u.2
-setSig (A:U) (B:A-> U) (sA: set A) (sB : (x:A) -> set (B x)) (t u : Sigma A B) : prop (Id (Sigma A B) t u) =
- substInv U prop (Id (Sigma A B) t u) ((p:T) * C p) rem3 rem2
+setSig (A:U) (B:A-> U) (sA: set A) (sB : (x:A) -> set (B x)) (t u : Sigma A B) : prop (Path (Sigma A B) t u) =
+ substInv U prop (Path (Sigma A B) t u) ((p:T) * C p) rem3 rem2
where
- T : U = Id A t.1 u.1
- C (p:T) : U = IdP (<i> B (p@i)) t.2 u.2
+ T : U = Path A t.1 u.1
+ C (p:T) : U = PathP (<i> B (p@i)) t.2 u.2
rem (p : T) : prop (C p) = corSigSet A B sB t u p
rem1 : prop T = sA t.1 u.1
rem2 : prop ((p:T) * C p) = propSig T C rem1 rem
- rem3 : Id U (Id (Sigma A B) t u) ((p:T) * C p) = lemIdSig A B t u
-
-corSigGroupoid (A:U) (B:A-> U) (gB : (x:A) -> groupoid (B x)) (t u : Sigma A B) (p:Id A t.1 u.1) :
- set (IdP (<i>B (p@i)) t.2 u.2) = substInv U set T0 T1 rem rem1
- where P : Id U (B t.1) (B u.1) = <i>B (p@i)
- T0 : U = IdP P t.2 u.2
- T1 : U = Id (B u.1) (transport P t.2) u.2
- rem : Id U T0 T1 = lem2 A B t u p -- funDepTr (B t.1) (B u.1) P t.2 u.2
+ rem3 : Path U (Path (Sigma A B) t u) ((p:T) * C p) = lemPathSig A B t u
+
+corSigGroupoid (A:U) (B:A-> U) (gB : (x:A) -> groupoid (B x)) (t u : Sigma A B) (p:Path A t.1 u.1) :
+ set (PathP (<i>B (p@i)) t.2 u.2) = substInv U set T0 T1 rem rem1
+ where P : Path U (B t.1) (B u.1) = <i>B (p@i)
+ T0 : U = PathP P t.2 u.2
+ T1 : U = Path (B u.1) (transport P t.2) u.2
+ rem : Path U T0 T1 = lem2 A B t u p -- funDepTr (B t.1) (B u.1) P t.2 u.2
v2 : B u.1 = transport P t.2
rem1 : set T1 = gB u.1 v2 u.2
-groupoidSig (A:U) (B:A-> U) (gA: groupoid A) (gB : (x:A) -> groupoid (B x)) (t u : Sigma A B) : set (Id (Sigma A B) t u) =
- substInv U set (Id (Sigma A B) t u) ((p:T) * C p) rem3 rem2
+groupoidSig (A:U) (B:A-> U) (gA: groupoid A) (gB : (x:A) -> groupoid (B x)) (t u : Sigma A B) : set (Path (Sigma A B) t u) =
+ substInv U set (Path (Sigma A B) t u) ((p:T) * C p) rem3 rem2
where
- T : U = Id A t.1 u.1
- C (p:T) : U = IdP (<i> B (p@i)) t.2 u.2
+ T : U = Path A t.1 u.1
+ C (p:T) : U = PathP (<i> B (p@i)) t.2 u.2
rem (p : T) : set (C p) = corSigGroupoid A B gB t u p
rem1 : set T = gA t.1 u.1
rem2 : set ((p:T) * C p) = setSig T C rem1 rem
- rem3 : Id U (Id (Sigma A B) t u) ((p:T) * C p) = lemIdSig A B t u
+ rem3 : Path U (Path (Sigma A B) t u) ((p:T) * C p) = lemPathSig A B t u
lemContr (A:U) (pA:prop A) (a:A) : isContr A = (a,rem)
- where rem (y:A) : Id A a y = pA a y
-
-lem3 (A:U) (B:A-> U) (pB : (x:A) -> prop (B x)) (t u : Sigma A B) (p:Id A t.1 u.1) :
- isContr (IdP (<i>B (p@i)) t.2 u.2) = lemContr T0 (substInv U prop T0 T1 rem rem1) rem2
- where P : Id U (B t.1) (B u.1) = <i>B (p@i)
- T0 : U = IdP P t.2 u.2
- T1 : U = Id (B u.1) (transport P t.2) u.2
- rem : Id U T0 T1 = lem2 A B t u p
+ where rem (y:A) : Path A a y = pA a y
+
+lem3 (A:U) (B:A-> U) (pB : (x:A) -> prop (B x)) (t u : Sigma A B) (p:Path A t.1 u.1) :
+ isContr (PathP (<i>B (p@i)) t.2 u.2) = lemContr T0 (substInv U prop T0 T1 rem rem1) rem2
+ where P : Path U (B t.1) (B u.1) = <i>B (p@i)
+ T0 : U = PathP P t.2 u.2
+ T1 : U = Path (B u.1) (transport P t.2) u.2
+ rem : Path U T0 T1 = lem2 A B t u p
v2 : B u.1 = transport P t.2
rem1 : prop T1 = propSet (B u.1) (pB u.1) v2 u.2
rem2 : T0 = transport (<i>rem@-i) (pB u.1 v2 u.2)
-lem6 (A:U) (P:A-> U) (cA:(x:A) -> isContr (P x)) : Id U ((x:A)*P x) A = isoId T A f g t s
+lem6 (A:U) (P:A-> U) (cA:(x:A) -> isContr (P x)) : Path U ((x:A)*P x) A = isoPath T A f g t s
where
T : U = (x:A) * P x
f (z:T) : A = z.1
g (x:A) : T = (x,(cA x).1)
- s (z:T) : Id T (g (f z)) z = <i>(z.1,((cA z.1).2 z.2)@ i)
- t (x:A) : Id A (f (g x)) x = refl A x
+ s (z:T) : Path T (g (f z)) z = <i>(z.1,((cA z.1).2 z.2)@ i)
+ t (x:A) : Path A (f (g x)) x = refl A x
-lemSigProp (A:U) (B:A-> U) (pB : (x:A) -> prop (B x)) (t u : Sigma A B) : Id U (Id (Sigma A B) t u) (Id A t.1 u.1) =
- compId U (Id (Sigma A B) t u) ((p:Id A t.1 u.1) * IdP (<i> B (p@i)) t.2 u.2) (Id A t.1 u.1) rem2 rem1
+lemSigProp (A:U) (B:A-> U) (pB : (x:A) -> prop (B x)) (t u : Sigma A B) : Path U (Path (Sigma A B) t u) (Path A t.1 u.1) =
+ compPath U (Path (Sigma A B) t u) ((p:Path A t.1 u.1) * PathP (<i> B (p@i)) t.2 u.2) (Path A t.1 u.1) rem2 rem1
where
- T : U = Id A t.1 u.1
- C (p:T) : U = IdP (<i> B (p@i)) t.2 u.2
+ T : U = Path A t.1 u.1
+ C (p:T) : U = PathP (<i> B (p@i)) t.2 u.2
rem (p : T) : isContr (C p) = lem3 A B pB t u p
- rem1 : Id U ((p:T) * C p) T = lem6 T C rem
- rem2 : Id U (Id (Sigma A B) t u) ((p:T) * C p) = lemIdSig A B t u
+ rem1 : Path U ((p:T) * C p) T = lem6 T C rem
+ rem2 : Path U (Path (Sigma A B) t u) ((p:T) * C p) = lemPathSig A B t u
-setGroupoid (A:U) (sA:set A) (a0 a1:A) : set (Id A a0 a1) = propSet (Id A a0 a1) (sA a0 a1)
+setGroupoid (A:U) (sA:set A) (a0 a1:A) : set (Path A a0 a1) = propSet (Path A a0 a1) (sA a0 a1)
propGroupoid (A:U) (pA: prop A) : groupoid A = setGroupoid A (propSet A pA)
= s.2.2.2
-- Construct a proposition to tag the element a with
X : U
- = (b : B) * (Id A (f b) a)
+ = (b : B) * (Path A (f b) a)
pX : prop X
= i a
in
(X, pX)
-lem (A:U) (P:A->U) (pP:(x:A) -> prop (P x)) (u v:(x:A) * P x) (p:Id A u.1 v.1) :
- Id ((x:A)*P x) u v = <i>(p@i,(lemPropF A P pP u.1 v.1 p u.2 v.2)@i)
+lem (A:U) (P:A->U) (pP:(x:A) -> prop (P x)) (u v:(x:A) * P x) (p:Path A u.1 v.1) :
+ Path ((x:A)*P x) u v = <i>(p@i,(lemPropF A P pP u.1 v.1 p u.2 v.2)@i)
-- A map from the second to the first definition of subsets.
subset10 (A : U) (sA : set A)
= \ (b : B) -> b.1
-- Show that f is injective.
inj : inj1 B A f sB sA
- = \ (a : A) (c d : (b : B) * Id A (f b) a) -> let
- p : Id A c.1.1 d.1.1
+ = \ (a : A) (c d : (b : B) * Path A (f b) a) -> let
+ p : Path A c.1.1 d.1.1
= <i> comp (<j> A) (c.2 @ i) [ (i = 0) -> <j> c.1.1
, (i = 1) -> <j> d.2 @ -j ]
- q : Id B c.1 d.1
+ q : Path B c.1 d.1
= lem A Q (\ (x : A) -> (P x).2) c.1 d.1 p
- r : Id ((b : B) * Id A (f b) a) c d
- = lem B (\(b : B) -> Id A (f b) a) (\ (b : B) -> sA (f b) a) c d q
+ r : Path ((b : B) * Path A (f b) a) c d
+ = lem B (\(b : B) -> Path A (f b) a) (\ (b : B) -> sA (f b) a) c d q
in
r
in
-- Show that subset10 ∘ subset01 can be identified with the identity function
subsetIso0 (A : U) (sA : set A) : (s0 : subset0 A sA) ->
- Id (subset0 A sA) (subset10 A sA (subset01 A sA s0)) s0
+ Path (subset0 A sA) (subset10 A sA (subset01 A sA s0)) s0
= \ (s0 : subset0 A sA) -> let
s0'
: subset0 A sA
= (f b, b, <i> f b)
g' (b' : B') : B
= b'.2.1
- s (x : B) : Id B (g' (g x)) x
+ s (x : B) : Path B (g' (g x)) x
= <i> x
- t (x : B') : Id B' (g (g' x)) x
+ t (x : B') : Path B' (g (g' x)) x
= <i> (x.2.2 @ i, x.2.1, <j> x.2.2 @ i /\ j)
-- Compute a path between B' and B, as well as a path between f'∘g∘g' and f
P (X : U) (h: X -> B) : U
- = (p : Id U X B) * (IdP (<i> p @ i -> A) (\ (x : X) -> f' (g (h x))) f)
+ = (p : Path U X B) * (PathP (<i> p @ i -> A) (\ (x : X) -> f' (g (h x))) f)
q : P B' g'
= elimEquiv B P (<i> B, <i> f) B' (g', gradLemma B' B g' g s t)
- idB : Id U B' B
+ idB : Path U B' B
= q.1
-- Show that sB can be identified with sB'
- idsB : IdP (<i> set (idB @ i)) sB' sB
+ idsB : PathP (<i> set (idB @ i)) sB' sB
= lemPropF U set setIsProp B' B idB sB' sB
-- Show that f' can be identified with f. This follows from g∘g' ⇔ \x -> x
-- and that there is a path q.2 between f'∘g∘g' and f
- idf : IdP (<i> idB @ i -> A) f' f
+ idf : PathP (<i> idB @ i -> A) f' f
= let
Q (h : B' -> B') : U
- = IdP (<i> q.1 @ i -> A) (\ (x : B') -> f' (h x)) f
+ = PathP (<i> q.1 @ i -> A) (\ (x : B') -> f' (h x)) f
a : B' -> B' = \ (x : B') -> g (g' x)
b : B' -> B' = \ (x : B') -> x
- p : Id (B' -> B') a b = <i> \ (x : B') -> (t x) @ i
+ p : Path (B' -> B') a b = <i> \ (x : B') -> (t x) @ i
in subst (B' -> B') Q a b p q.2
-- Show that inj can be identified with inj'
- idinj : IdP (<i> inj1 (idB @ i) A (idf @ i) (idsB @ i) sA) inj' inj
+ idinj : PathP (<i> inj1 (idB @ i) A (idf @ i) (idsB @ i) sA) inj' inj
= let
T : U
= (X : U) * (_ : X -> A) * (set X)
= (B', f', sB')
t1 : T
= (B, f, sB)
- idT : Id T t0 t1
+ idT : Path T t0 t1
= <i> (idB @ i, idf @ i, idsB @ i)
in lemPropF T P pP t0 t1 idT inj' inj
in
-- Show that subset10 ∘ subset01 can be identified with the identity function
subsetIso1 (A : U) (sA : set A) : (s1 : subset1 A sA) ->
- Id (subset1 A sA) (subset01 A sA (subset10 A sA s1)) s1
+ Path (subset1 A sA) (subset01 A sA (subset10 A sA s1)) s1
= \ (s1 : subset1 A sA) -> let
-- Construct the second subset s1' from s1.
s1' : subset1 A sA
= subset01 A sA (subset10 A sA s1)
-- Show that s1' and s1 produces the same result for all a : A
- ids1 : (a : A) -> Id ((X : U) * (prop X)) (s1' a) (s1 a)
+ ids1 : (a : A) -> Path ((X : U) * (prop X)) (s1' a) (s1 a)
= \ (a : A) -> let
-- Construct isomorphism between (s1' a).1 and (s1 a).1 to show that
-- (s1' a).1 can be identified with (s1 a).1
= \ (x : (s1' a).1) -> subst A (\(a : A) -> (s1 a).1) x.1.1 a x.2 x.1.2
g : (s1 a).1 -> (s1' a).1
= \ (x : (s1 a).1) -> ((a, x), <i> a)
- s : (x : (s1 a).1) -> Id (s1 a).1 (f (g x)) x
+ s : (x : (s1 a).1) -> Path (s1 a).1 (f (g x)) x
= \ (x : (s1 a).1) -> (s1 a).2 (f (g x)) x
- t : (x : (s1' a).1) -> Id (s1' a).1 (g (f x)) x
+ t : (x : (s1' a).1) -> Path (s1' a).1 (g (f x)) x
= \ (x : (s1' a).1) -> (s1' a).2 (g (f x)) x
- p : Id U (s1' a).1 (s1 a).1
- = isoId (s1' a).1 (s1 a).1 f g s t
+ p : Path U (s1' a).1 (s1 a).1
+ = isoPath (s1' a).1 (s1 a).1 f g s t
-- Show that for x : prop (s1' a).1, y : prop (s1 a).1,
-- x can be identified with y.
- q : IdP (<i> prop (p @ i)) (s1' a).2 (s1 a).2
+ q : PathP (<i> prop (p @ i)) (s1' a).2 (s1 a).2
= lemPropF U prop propIsProp (s1' a).1 (s1 a).1 p (s1' a).2 (s1 a).2
in
<i> (p @ i, q @ i)
funExt A (\ (_ : A) -> (X : U) * (prop X)) s1' s1 ids1
-- Show that we can identify the two definitions of subsets with each other
-subsetId (A : U) (sA : set A) : Id U (subset0 A sA) (subset1 A sA)
- = isoId (subset0 A sA) (subset1 A sA) (subset01 A sA) (subset10 A sA)
+subsetPath (A : U) (sA : set A) : Path U (subset0 A sA) (subset1 A sA)
+ = isoPath (subset0 A sA) (subset1 A sA) (subset01 A sA) (subset10 A sA)
(subsetIso1 A sA) (subsetIso0 A sA)
\ No newline at end of file
-- (Similar to HoTT Book, Lemma 6.5.1)\r
sone : U = sphere one\r
\r
-path : bool -> Id S1 base base = split\r
+path : bool -> Path S1 base base = split\r
false -> loop1\r
true -> refl S1 base\r
\r
south -> base\r
merid b @ i -> path b @ i\r
\r
-m0 : Id sone north south = <i> merid{sone} false @ i\r
+m0 : Path sone north south = <i> merid{sone} false @ i\r
\r
-m1 : Id sone north south = <i> merid{sone} true @ i\r
+m1 : Path sone north south = <i> merid{sone} true @ i\r
\r
-invm1 : Id sone south north = inv sone north south m1\r
+invm1 : Path sone south north = inv sone north south m1\r
\r
circleToS1 : S1 -> sone = split\r
base -> north\r
- loop @ i -> compId sone north south north m0 invm1 @ i\r
+ loop @ i -> compPath sone north south north m0 invm1 @ i\r
\r
-merid1 (b:bool) : Id sone north south = <i> merid{sone} b @ i\r
+merid1 (b:bool) : Path sone north south = <i> merid{sone} b @ i\r
\r
co (x: sone) : sone = circleToS1 (s1ToCircle x)\r
\r
-lemSquare (A:U) (a b : A) (m0 m1 : Id A a b) :\r
- Square A a a a b (compId A a b a m0 (inv A a b m1)) m0 (refl A a) m1 =\r
+lemSquare (A:U) (a b : A) (m0 m1 : Path A a b) :\r
+ Square A a a a b (compPath A a b a m0 (inv A a b m1)) m0 (refl A a) m1 =\r
<i j> comp (<_>A) (m0 @ i) [(i=1) -> <k> m1 @ (j \/ -k),\r
(i=0) -> <_>a,\r
(j=1) -> <_>m0@i,\r
(j=0) -> <k> comp (<_>A) (m0 @ i) [(k=0) -> <_>m0@i, (i=0) -> <_>a, (i=1) -> <l> m1 @ (-k \/ -l)]]\r
\r
-coid : (x : sone) -> Id sone (co x) x = split\r
+coid : (x : sone) -> Path sone (co x) x = split\r
north -> refl sone north\r
south -> m1\r
merid b @ i -> ind b @ i\r
where\r
- F (x:sone) : U = Id sone (co x) x\r
+ F (x:sone) : U = Path sone (co x) x\r
\r
- ind : (b:bool) -> IdS sone F north south (merid1 b) (refl sone north) m1 = split\r
+ ind : (b:bool) -> PathS sone F north south (merid1 b) (refl sone north) m1 = split\r
false -> lemSquare sone north south m0 m1\r
true -> <j k> m1 @ (j /\ k)\r
\r
oc (x:S1) : S1 = s1ToCircle (circleToS1 x)\r
\r
-ocid : (x : S1) -> Id S1 (oc x) x = \r
+ocid : (x : S1) -> Path S1 (oc x) x = \r
split\r
base -> refl S1 base\r
loop @ i -> <j>comp (<_>S1) (loop1@i) [(i=0) -> <_>base,(i=1) -> <_>base,(j=1) -> <_>loop1@i, \r
\r
\r
\r
-s1EqCircle : Id U sone S1 = isoId sone S1 s1ToCircle circleToS1 ocid coid\r
+s1EqCircle : Path U sone S1 = isoPath sone S1 s1ToCircle circleToS1 ocid coid\r
\r
-s1EqS1 : Id U S1 S1 = compId U S1 sone S1 (inv U sone S1 s1EqCircle) s1EqCircle\r
+s1EqS1 : Path U S1 S1 = compPath U S1 sone S1 (inv U sone S1 s1EqCircle) s1EqCircle\r
\r
-lem (A:U) (a:A) : Id A (comp (<_>A) (comp (<_>A) (comp (<_>A) a []) []) []) a = \r
+lem (A:U) (a:A) : Path A (comp (<_>A) (comp (<_>A) (comp (<_>A) a []) []) []) a = \r
<i>comp (<_>A) (comp (<_>A) (comp (<_>A) a [(i=1) -> <_>a]) [(i=1) -> <_>a]) [(i=1) -> <_>a]\r
\r
\r
\r
ptU : U = (X : U) * X\r
\r
-lemPt (A :U) (B:U) (p:Id U A B) (a:A) : Id ptU (A,a) (B,transport p a) =\r
+lemPt (A :U) (B:U) (p:Path U A B) (a:A) : Path ptU (A,a) (B,transport p a) =\r
<i> (p @ i,comp (<j> p @ (i/\j)) a [(i=0) -> <_>a])\r
\r
-Omega (X:ptU) : ptU = (Id X.1 X.2 X.2,refl X.1 X.2)\r
+Omega (X:ptU) : ptU = (Path X.1 X.2 X.2,refl X.1 X.2)\r
\r
-lem (A:U) (a:A) : Id A (comp (<_>A) (comp (<_>A) (comp (<_>A) a []) []) []) a = \r
+lem (A:U) (a:A) : Path A (comp (<_>A) (comp (<_>A) (comp (<_>A) a []) []) []) a = \r
<i>comp (<_>A) (comp (<_>A) (comp (<_>A) a [(i=1) -> <_>a]) [(i=1) -> <_>a]) [(i=1) -> <_>a]\r
\r
-lem1 (A:U) (a:A) : Id ptU (A,comp (<_>A) (comp (<_>A) (comp (<_>A) a []) []) []) (A,a) = \r
+lem1 (A:U) (a:A) : Path ptU (A,comp (<_>A) (comp (<_>A) (comp (<_>A) a []) []) []) (A,a) = \r
<i>(A,lem A a@i)\r
\r
--- s1PtCircle : Id ptU (sone,north) (S1,base) = \r
--- compId ptU (sone,north) (S1,comp (<_>S1) (comp (<_>S1) (comp (<_>S1) base []) []) []) (S1,base) (lemPt sone S1 s1EqCircle north) (lem1 S1 base) \r
+-- s1PtCircle : Path ptU (sone,north) (S1,base) = \r
+-- compPath ptU (sone,north) (S1,comp (<_>S1) (comp (<_>S1) (comp (<_>S1) base []) []) []) (S1,base) (lemPt sone S1 s1EqCircle north) (lem1 S1 base) \r
\r
--- windingS : Id sone north north -> Z = rem1\r
+-- windingS : Path sone north north -> Z = rem1\r
-- where\r
-- G (X:ptU) : U = (Omega X).1 -> Z\r
-- rem : G (S1,base) = winding\r
-- rem1 : G (sone,north) = subst ptU G (S1,base) (sone,north) (<i> s1PtCircle @ -i) rem\r
\r
--- s1ToCId (p: Id sone north north) : Id S1 base base = <i> transport s1EqCircle (p @ i)\r
+-- s1ToCPath (p: Path sone north north) : Path S1 base base = <i> transport s1EqCircle (p @ i)\r
\r
--- s1ToCIdInv (p : Id S1 base base) : Id sone north north = <i> (transport (<j> s1EqCircle @ -j) (p @ i))\r
+-- s1ToCPathInv (p : Path S1 base base) : Path sone north north = <i> (transport (<j> s1EqCircle @ -j) (p @ i))\r
\r
-loop1S : Id sone north north = compId sone north south north m0 invm1\r
+loop1S : Path sone north north = compPath sone north south north m0 invm1\r
\r
-loop2S : Id sone north north = compId sone north north north loop1S loop1S\r
+loop2S : Path sone north north = compPath sone north north north loop1S loop1S\r
\r
-- test0S : Z = windingS (refl sone north)\r
\r
-- test2S : Z = windingS loop2S\r
\r
--- test4S : Z = windingS (compId sone north north north loop2S loop2S)\r
+-- test4S : Z = windingS (compPath sone north north north loop2S loop2S)\r
\r
\r
-- indSusp:\r
\r
\r
-suspOf (A X : U) : U = (u:X) * (v:X) * (A -> Id X u v)\r
+suspOf (A X : U) : U = (u:X) * (v:X) * (A -> Path X u v)\r
\r
funToL (A X:U) (f:susp A -> X) : suspOf A X =\r
(f north,f south,\ (a:A) -> <i>f (merid{susp A} a@i))\r
south -> z.2.1\r
merid a @ i-> z.2.2 a @ i\r
\r
-test1 (A X:U) (z:suspOf A X) : Id (suspOf A X) (funToL A X (lToFun A X z)) z\r
+test1 (A X:U) (z:suspOf A X) : Path (suspOf A X) (funToL A X (lToFun A X z)) z\r
= refl (suspOf A X) z\r
\r
-rem (A X:U) (f:susp A ->X) : (u:susp A) -> Id X (lToFun A X (funToL A X f) u) (f u) = split\r
+rem (A X:U) (f:susp A ->X) : (u:susp A) -> Path X (lToFun A X (funToL A X f) u) (f u) = split\r
north -> refl X (f north)\r
south -> refl X (f south)\r
merid a @ i -> refl X (f (merid{susp A} a @ i))\r
\r
-test2 (A X:U) (f:susp A ->X) : Id (susp A ->X) (lToFun A X (funToL A X f)) f\r
+test2 (A X:U) (f:susp A ->X) : Path (susp A ->X) (lToFun A X (funToL A X f)) f\r
= <i>\ (u:susp A) -> rem A X f u @ i\r
\r
-funSusp (A X:U) : Id U (susp A -> X) (suspOf A X) =\r
- isoId (susp A -> X) (suspOf A X) (funToL A X) (lToFun A X) (test1 A X) (test2 A X)\r
+funSusp (A X:U) : Path U (susp A -> X) (suspOf A X) =\r
+ isoPath (susp A -> X) (suspOf A X) (funToL A X) (lToFun A X) (test1 A X) (test2 A X)\r
isEquivComp (A B C : U) (F : A -> B) (G : B -> C) (ef : isEquiv A B F) (eg : isEquiv B C G)\r
: isEquiv A C (\(x : A) -> G (F x))\r
= gradLemma A C (\(x : A) -> G (F x)) (\(x : C) -> (ef (eg x).1.1).1.1)\r
- (\(x : C) -> compId C (G (F (ef (eg x).1.1).1.1)) (G (eg x).1.1) x\r
+ (\(x : C) -> compPath C (G (F (ef (eg x).1.1).1.1)) (G (eg x).1.1) x\r
(<i> G (retEq A B (F, ef) (eg x).1.1 @ i)) (retEq B C (G, eg) x))\r
- (\(x : A) -> compId A ((ef (eg (G (F x))).1.1).1.1) (ef (F x)).1.1 x\r
+ (\(x : A) -> compPath A ((ef (eg (G (F x))).1.1).1.1) (ef (F x)).1.1 x\r
(<i> (ef (secEq B C (G, eg) (F x) @ i)).1.1) (secEq A B (F, ef) x))\r
\r
isEquivComp' (A B C : U) (F : A -> B) (G : C -> B) (ef : isEquiv A B F) (eg : isEquiv C B G)\r
: isEquiv A C (\(x : A) -> (eg (F x)).1.1)\r
= gradLemma A C (\(x : A) -> (eg (F x)).1.1) (\(x : C) -> (ef (G x)).1.1)\r
- (\(x : C) -> compId C (eg (F (ef (G x)).1.1)).1.1 (eg (G x)).1.1 x\r
+ (\(x : C) -> compPath C (eg (F (ef (G x)).1.1)).1.1 (eg (G x)).1.1 x\r
(<i> (eg (retEq A B (F, ef) (G x) @ i)).1.1) (secEq C B (G, eg) x))\r
- (\(x : A) -> compId A ((ef (G (eg (F x)).1.1)).1.1) (ef (F x)).1.1 x\r
+ (\(x : A) -> compPath A ((ef (G (eg (F x)).1.1)).1.1) (ef (F x)).1.1 x\r
(<i> (ef (retEq C B (G, eg) (F x) @ i)).1.1) (secEq A B (F, ef) x))\r
\r
\r
-lemHcomp (x : loopS1) : Id loopS1 (<i> comp (<_>S1) (x@i) [(i=0)-><_>base,(i=1)-><_>base]) x\r
+lemHcomp (x : loopS1) : Path loopS1 (<i> comp (<_>S1) (x@i) [(i=0)-><_>base,(i=1)-><_>base]) x\r
= <j i> comp (<_>S1) (x@i) [(i=0)-><_>base,(j=1)-><_>x@i,(i=1)-><_>base]\r
opaque lemHcomp\r
\r
lemHcomp3 (x : loopS1)\r
- : Id loopS1\r
+ : Path loopS1\r
(<i> comp (<_>S1) (comp (<_>S1) (comp (<_>S1) (x@i) [(i=0)-><_>base,(i=1)-><_>base]) [(i=0)-><_>base,(i=1)-><_>base]) [(i=0)-><_>base,(i=1)-><_>base]) x\r
- = compId loopS1\r
+ = compPath loopS1\r
(<i> comp (<_>S1) (comp (<_>S1) (comp (<_>S1) (x@i) [(i=0)-><_>base,(i=1)-><_>base]) [(i=0)-><_>base,(i=1)-><_>base]) [(i=0)-><_>base,(i=1)-><_>base])\r
(<i> comp (<_>S1) (comp (<_>S1) (x@i) [(i=0)-><_>base,(i=1)-><_>base]) [(i=0)-><_>base,(i=1)-><_>base])\r
x\r
(<j i> comp (<_>S1) (comp (<_>S1) (lemHcomp x@j@i) [(i=0)-><_>base,(i=1)-><_>base]) [(i=0)-><_>base,(i=1)-><_>base])\r
- (compId loopS1\r
+ (compPath loopS1\r
(<i> comp (<_>S1) (comp (<_>S1) (x@i) [(i=0)-><_>base,(i=1)-><_>base]) [(i=0)-><_>base,(i=1)-><_>base])\r
(<i> comp (<_>S1) (x@i) [(i=0)-><_>base,(i=1)-><_>base])\r
x\r
(lemHcomp x))\r
opaque lemHcomp3\r
\r
-lemEquiv1 (A B : U) (F : A -> B) (G : A -> B) (e : isEquiv A B G) (p : (x : A) -> Id A (e (F x)).1.1 x) : Id (A -> B) F G\r
- = <i> \(x : A) -> transport (<i> Id B (retEq A B (G, e) (F x) @ i) (G x)) (mapOnPath A B G (e (F x)).1.1 x (p x)) @ i\r
+lemEquiv1 (A B : U) (F : A -> B) (G : A -> B) (e : isEquiv A B G) (p : (x : A) -> Path A (e (F x)).1.1 x) : Path (A -> B) F G\r
+ = <i> \(x : A) -> transport (<i> Path B (retEq A B (G, e) (F x) @ i) (G x)) (mapOnPath A B G (e (F x)).1.1 x (p x)) @ i\r
\r
-transRefl (A : U) (a : A) : Id A (transport (<_> A) a) a = <i> comp (<_> A) a [(i=1) -> <_>a]\r
-lemReflComp (A : U) (a b : A) (p : Id A a b) : Id (Id A a b) (compId A a a b (<_> a) p) p =\r
+transRefl (A : U) (a : A) : Path A (transport (<_> A) a) a = <i> comp (<_> A) a [(i=1) -> <_>a]\r
+lemReflComp (A : U) (a b : A) (p : Path A a b) : Path (Path A a b) (compPath A a a b (<_> a) p) p =\r
<j i> comp (<k> A) (p @ i /\ j) [(i=0) -> <_> a, (j=1) -> <_> p @ i, (i=1) -> <k> p @ k \/ j ]\r
-lemReflComp' (A : U) (a b : A) (p : Id A a b) : Id (Id A a b) (compId A a b b p (<_> b)) p =\r
+lemReflComp' (A : U) (a b : A) (p : Path A a b) : Path (Path A a b) (compPath A a b b p (<_> b)) p =\r
<j i> comp (<k> A) (p @ i) [(i=0) -> <_> a, (j=1) -> <_> p @ i, (i=1) -> <_> b ]\r
\r
-lem2ItPos : (n:nat) -> Id loopS1 (loopIt (predZ (inr n))) (backTurn (loopIt (inr n))) = split\r
- zero -> transport (<i> Id loopS1 invLoop (lemReflComp S1 base base invLoop @ -i)) (<_> invLoop)\r
+lem2ItPos : (n:nat) -> Path loopS1 (loopIt (predZ (inr n))) (backTurn (loopIt (inr n))) = split\r
+ zero -> transport (<i> Path loopS1 invLoop (lemReflComp S1 base base invLoop @ -i)) (<_> invLoop)\r
suc p -> compInv S1 base base (loopIt (inr p)) base loop1\r
\r
-lem2It : (n:Z) -> Id loopS1 (loopIt (predZ n)) (backTurn (loopIt n)) = split\r
+lem2It : (n:Z) -> Path loopS1 (loopIt (predZ n)) (backTurn (loopIt n)) = split\r
inl n -> <_> loopIt (inl (suc n))\r
inr n -> lem2ItPos n\r
\r
-transportCompId (a b c : U) (p1 : Id U a b) (p2 : Id U b c) (x : a)\r
- : Id c (transport (compId U a b c p1 p2) x) (transport p2 (transport p1 x))\r
- = J U b (\(c : U) -> \(p2 : Id U b c) -> Id c (comp (compId U a b c p1 p2) x []) (comp p2 (transport p1 x) []))\r
+transportCompPath (a b c : U) (p1 : Path U a b) (p2 : Path U b c) (x : a)\r
+ : Path c (transport (compPath U a b c p1 p2) x) (transport p2 (transport p1 x))\r
+ = J U b (\(c : U) -> \(p2 : Path U b c) -> Path c (comp (compPath U a b c p1 p2) x []) (comp p2 (transport p1 x) []))\r
hole c p2\r
where\r
- hole : Id b (comp (compId U a b b p1 (<_> b)) x []) (comp (<_> b) (transport p1 x) []) =\r
- compId b (comp (compId U a b b p1 (<_> b)) x []) (transport p1 x) (comp (<_> b) (transport p1 x) [])\r
+ hole : Path b (comp (compPath U a b b p1 (<_> b)) x []) (comp (<_> b) (transport p1 x) []) =\r
+ compPath b (comp (compPath U a b b p1 (<_> b)) x []) (transport p1 x) (comp (<_> b) (transport p1 x) [])\r
(<i> comp (lemReflComp' U a b p1 @ i) x [])\r
(<i> transRefl b (transport p1 x) @ -i)\r
\r
-lemRevCompId (A : U) (a b c : A) (p1 : Id A a b) (p2 : Id A b c)\r
- : Id (Id A c a) (<i> compId A a b c p1 p2 @ -i) (compId A c b a (<i> p2@-i) (<i> p1@-i))\r
+lemRevCompPath (A : U) (a b c : A) (p1 : Path A a b) (p2 : Path A b c)\r
+ : Path (Path A c a) (<i> compPath A a b c p1 p2 @ -i) (compPath A c b a (<i> p2@-i) (<i> p1@-i))\r
= <j i> comp (<k> A) (hole1 @ i @ j) [(i=0) -> <k> p2 @ k \/ j, (i=1) -> <k> p1 @ -k /\ j]\r
where\r
hole1 : Square A b a c b (<i> p1@-i) (<i> p2@-i) p2 p1\r
\r
setPi (A : U) (B : A -> U) (h : (x : A) -> set (B x))\r
(f0 f1 : (x : A) -> B x)\r
- (p1 p2 : Id ((x : A) -> B x) f0 f1)\r
- : Id (Id ((x : A) -> B x) f0 f1) p1 p2\r
+ (p1 p2 : Path ((x : A) -> B x) f0 f1)\r
+ : Path (Path ((x : A) -> B x) f0 f1) p1 p2\r
= <i j> \(x : A) -> (h x (f0 x) (f1 x) (<i> (p1@i) x) (<i> (p2@i) x)) @ i @ j\r
\r
-lemIdPProp (A B : U) (AProp : prop A) (p : Id U A B) : (x : A) -> (y : B) -> IdP p x y\r
- = J U A (\(B : U) -> \(p : Id U A B) -> (x : A) -> (y : B) -> IdP p x y) AProp B p\r
+lemPathPProp (A B : U) (AProp : prop A) (p : Path U A B) : (x : A) -> (y : B) -> PathP p x y\r
+ = J U A (\(B : U) -> \(p : Path U A B) -> (x : A) -> (y : B) -> PathP p x y) AProp B p\r
\r
-lemIdPSet (A B : U) (ASet : set A) (p : Id U A B) : (x : A) (y : B) (s t : IdP p x y) -> Id (IdP p x y) s t\r
- = J U A (\(B : U) -> \(p : Id U A B) -> (x : A) (y : B) (s t : IdP p x y) -> Id (IdP p x y) s t) ASet B p\r
+lemPathPSet (A B : U) (ASet : set A) (p : Path U A B) : (x : A) (y : B) (s t : PathP p x y) -> Path (PathP p x y) s t\r
+ = J U A (\(B : U) -> \(p : Path U A B) -> (x : A) (y : B) (s t : PathP p x y) -> Path (PathP p x y) s t) ASet B p\r
\r
-lemIdPSet2 (A B : U) (ASet : set A) (p1 : Id U A B)\r
- : (p2 : Id U A B) -> (p : Id (Id U A B) p1 p2) ->\r
- (x : A) -> (y : B) -> (s : IdP p1 x y) -> (t : IdP p2 x y) -> IdP (<i> (IdP (p @ i) x y)) s t\r
- = J (Id U A B) p1 (\(p2 : Id U A B) -> \(p : Id (Id U A B) p1 p2) -> (x : A) -> (y : B) -> (s : IdP p1 x y) -> (t : IdP p2 x y) -> IdP (<i> (IdP (p @ i) x y)) s t)\r
- (lemIdPSet A B ASet p1)\r
+lemPathPSet2 (A B : U) (ASet : set A) (p1 : Path U A B)\r
+ : (p2 : Path U A B) -> (p : Path (Path U A B) p1 p2) ->\r
+ (x : A) -> (y : B) -> (s : PathP p1 x y) -> (t : PathP p2 x y) -> PathP (<i> (PathP (p @ i) x y)) s t\r
+ = J (Path U A B) p1 (\(p2 : Path U A B) -> \(p : Path (Path U A B) p1 p2) -> (x : A) -> (y : B) -> (s : PathP p1 x y) -> (t : PathP p2 x y) -> PathP (<i> (PathP (p @ i) x y)) s t)\r
+ (lemPathPSet A B ASet p1)\r
\r
\r
-isEquivId (A : U) : isEquiv A A (\(x : A) -> x) = gradLemma A A (\(x : A) -> x) (\(x : A) -> x) (\(x : A) -> <_> x) (\(x : A) -> <_> x)\r
+isEquivPath (A : U) : isEquiv A A (\(x : A) -> x) = gradLemma A A (\(x : A) -> x) (\(x : A) -> x) (\(x : A) -> <_> x) (\(x : A) -> <_> x)\r
\r
-lem10 (A B : U) (e : equiv A B) (x y : B) (p : Id A (e.2 x).1.1 (e.2 y).1.1) : Id B x y\r
+lem10 (A B : U) (e : equiv A B) (x y : B) (p : Path A (e.2 x).1.1 (e.2 y).1.1) : Path B x y\r
= transport\r
- (compId U (Id B (e.1 (e.2 x).1.1) (e.1 (e.2 y).1.1)) (Id B x (e.1 (e.2 y).1.1)) (Id B x y)\r
- (<i> Id B (retEq A B e x @ i) (e.1 (e.2 y).1.1)) (<i> Id B x (retEq A B e y @ i)))\r
+ (compPath U (Path B (e.1 (e.2 x).1.1) (e.1 (e.2 y).1.1)) (Path B x (e.1 (e.2 y).1.1)) (Path B x y)\r
+ (<i> Path B (retEq A B e x @ i) (e.1 (e.2 y).1.1)) (<i> Path B x (retEq A B e y @ i)))\r
(mapOnPath A B e.1 (e.2 x).1.1 (e.2 y).1.1 p)\r
\r
-lem10' (A B : U) (e : equiv A B) (x y : A) (p : Id B (e.1 x) (e.1 y)) : Id A x y\r
+lem10' (A B : U) (e : equiv A B) (x y : A) (p : Path B (e.1 x) (e.1 y)) : Path A x y\r
= transport\r
- (compId U (Id A (e.2 (e.1 x)).1.1 (e.2 (e.1 y)).1.1) (Id A x (e.2 (e.1 y)).1.1) (Id A x y)\r
- (<i> Id A (secEq A B e x @ i) (e.2 (e.1 y)).1.1) (<i> Id A x (secEq A B e y @ i))\r
+ (compPath U (Path A (e.2 (e.1 x)).1.1 (e.2 (e.1 y)).1.1) (Path A x (e.2 (e.1 y)).1.1) (Path A x y)\r
+ (<i> Path A (secEq A B e x @ i) (e.2 (e.1 y)).1.1) (<i> Path A x (secEq A B e y @ i))\r
)\r
(mapOnPath B A (invEq A B e) (e.1 x) (e.1 y) p)\r
\r
-lem11 (A : U) : Id (Id U A A) (univalence A A (\(x : A) -> x, isEquivId A)).1.1 (<_> A) = hole\r
+lem11 (A : U) : Path (Path U A A) (univalence A A (\(x : A) -> x, isEquivPath A)).1.1 (<_> A) = hole\r
where\r
- hole0 : Id (equiv A A)\r
- (\(x : A) -> x, isEquivId A)\r
+ hole0 : Path (equiv A A)\r
+ (\(x : A) -> x, isEquivPath A)\r
(transEquiv' A A (<_> A))\r
= <i> ( (transRefl (A->A) (\(x : A) -> x) @ -i)\r
- , lemIdPProp (isEquiv A A (\(x : A) -> x)) (isEquiv A A (transEquiv' A A (<_> A)).1)\r
+ , lemPathPProp (isEquiv A A (\(x : A) -> x)) (isEquiv A A (transEquiv' A A (<_> A)).1)\r
(propIsEquiv A A (\(x : A) -> x)) (<j> isEquiv A A (transRefl (A->A) (\(x : A) -> x) @ -j))\r
- (isEquivId A) (transEquiv' A A (<_> A)).2 @ i\r
+ (isEquivPath A) (transEquiv' A A (<_> A)).2 @ i\r
)\r
- hole1 : Id (equiv A A) (transEquiv' A A (univalence A A (\(x : A) -> x, isEquivId A)).1.1) (\(x : A) -> x, isEquivId A)\r
- = retEq (Id U A A) (equiv A A) (corrUniv' A A) (\(x : A) -> x, isEquivId A)\r
- hole : Id (Id U A A) (univalence A A (\(x : A) -> x, isEquivId A)).1.1 (<_> A)\r
- = lem10' (Id U A A) (equiv A A) (corrUniv' A A)\r
- (univalence A A (\(x : A) -> x, isEquivId A)).1.1 (<_> A)\r
- (compId (equiv A A) (transEquiv' A A (univalence A A (\(x : A) -> x, isEquivId A)).1.1) (\(x : A) -> x, isEquivId A) (transEquiv' A A (<_> A)) hole1 hole0)\r
+ hole1 : Path (equiv A A) (transEquiv' A A (univalence A A (\(x : A) -> x, isEquivPath A)).1.1) (\(x : A) -> x, isEquivPath A)\r
+ = retEq (Path U A A) (equiv A A) (corrUniv' A A) (\(x : A) -> x, isEquivPath A)\r
+ hole : Path (Path U A A) (univalence A A (\(x : A) -> x, isEquivPath A)).1.1 (<_> A)\r
+ = lem10' (Path U A A) (equiv A A) (corrUniv' A A)\r
+ (univalence A A (\(x : A) -> x, isEquivPath A)).1.1 (<_> A)\r
+ (compPath (equiv A A) (transEquiv' A A (univalence A A (\(x : A) -> x, isEquivPath A)).1.1) (\(x : A) -> x, isEquivPath A) (transEquiv' A A (<_> A)) hole1 hole0)\r
opaque lem11\r
\r
-substIdP (A B : U) (p : Id U A B) (x : A) (y : B) (q : Id B (transport p x) y) : IdP p x y\r
- = transport (<i> IdP p x (q@i)) hole\r
+substPathP (A B : U) (p : Path U A B) (x : A) (y : B) (q : Path B (transport p x) y) : PathP p x y\r
+ = transport (<i> PathP p x (q@i)) hole\r
where\r
- hole : IdP p x (transport p x) = <i> comp (<j> p @ (i /\ j)) x [(i=0) -> <_> x]\r
+ hole : PathP p x (transport p x) = <i> comp (<j> p @ (i /\ j)) x [(i=0) -> <_> x]\r
\r
-lemIdSig (A:U) (B : A -> U) (t u : Sigma A B) :\r
- Id U (Id (Sigma A B) t u) ((p : Id A t.1 u.1) * IdP (<i> B (p @ i)) t.2 u.2) =\r
- isoId T0 T1 f g s t where\r
- T0 : U = Id (Sigma A B) t u\r
- T1 : U = (p:Id A t.1 u.1) * IdP (<i> B (p@i)) t.2 u.2\r
+lemPathSig (A:U) (B : A -> U) (t u : Sigma A B) :\r
+ Path U (Path (Sigma A B) t u) ((p : Path A t.1 u.1) * PathP (<i> B (p @ i)) t.2 u.2) =\r
+ isoPath T0 T1 f g s t where\r
+ T0 : U = Path (Sigma A B) t u\r
+ T1 : U = (p:Path A t.1 u.1) * PathP (<i> B (p@i)) t.2 u.2\r
f (q:T0) : T1 = (<i> (q@i).1,<i> (q@i).2)\r
g (z:T1) : T0 = <i>(z.1 @i,z.2 @i)\r
- s (z:T1) : Id T1 (f (g z)) z = refl T1 z\r
- t (q:T0) : Id T0 (g (f q)) q = refl T0 q\r
+ s (z:T1) : Path T1 (f (g z)) z = refl T1 z\r
+ t (q:T0) : Path T0 (g (f q)) q = refl T0 q\r
\r
transConstN (A : U) (a : A) : (n : nat) -> A = split\r
zero -> a\r
suc n -> transport (<_> A) (transConstN A a n)\r
\r
-transConstNRefl (A : U) (a : A) : (n : nat) -> Id A (transConstN A a n) a = split\r
+transConstNRefl (A : U) (a : A) : (n : nat) -> Path A (transConstN A a n) a = split\r
zero -> <_> a\r
- suc n -> compId A (transport (<_> A) (transConstN A a n)) (transConstN A a n) a\r
+ suc n -> compPath A (transport (<_> A) (transConstN A a n)) (transConstN A a n) a\r
(transRefl A (transConstN A a n)) (transConstNRefl A a n)\r
\r
lem0 (A B : U) (f : A -> B) (e : isEquiv A B f) (x : A)\r
- : Id B (transport (univalence B A (f, e)).1.1 x) (f x)\r
+ : Path B (transport (univalence B A (f, e)).1.1 x) (f x)\r
= transConstNRefl B (f x) (suc (suc (suc (suc (suc (suc (suc (suc (suc (suc (suc zero)))))))))))\r
\r
lem1 (A B : U) (f : A -> B) (e : isEquiv A B f) (x : B)\r
- : Id A (transport (<i> (univalence B A (f, e)).1.1 @ -i) x) (e x).1.1\r
- = compId A\r
+ : Path A (transport (<i> (univalence B A (f, e)).1.1 @ -i) x) (e x).1.1\r
+ = compPath A\r
(transConstN A (e (transConstN B x (suc (suc (suc (suc (suc (suc (suc (suc zero)))))))))).1.1 (suc (suc (suc zero))))\r
(transConstN A (e x).1.1 (suc (suc (suc zero))))\r
(e x).1.1\r
\r
-- P1 <-> P2 <-> P3\r
\r
-P1 (A B : U) (F : A -> B) : U = (x y : A) -> isEquiv (Id A x y) (Id B (F x) (F y)) (mapOnPath A B F x y)\r
-P2 (A B : U) (F : A -> B) : U = (x : A) -> isContr ((y : A) * Id B (F x) (F y))\r
-P3 (A B : U) (F : A -> B) : U = (x : B) -> prop ((y : A) * Id B x (F y))\r
+P1 (A B : U) (F : A -> B) : U = (x y : A) -> isEquiv (Path A x y) (Path B (F x) (F y)) (mapOnPath A B F x y)\r
+P2 (A B : U) (F : A -> B) : U = (x : A) -> isContr ((y : A) * Path B (F x) (F y))\r
+P3 (A B : U) (F : A -> B) : U = (x : B) -> prop ((y : A) * Path B x (F y))\r
\r
propP1 (A B : U) (F : A -> B) : prop (P1 A B F)\r
- = propPi A (\(x : A) -> (y : A) -> isEquiv (Id A x y) (Id B (F x) (F y)) (mapOnPath A B F x y))\r
- (\(x : A) -> propPi A (\(y : A) -> isEquiv (Id A x y) (Id B (F x) (F y)) (mapOnPath A B F x y))\r
- (\(y : A) -> propIsEquiv (Id A x y) (Id B (F x) (F y)) (mapOnPath A B F x y)))\r
+ = propPi A (\(x : A) -> (y : A) -> isEquiv (Path A x y) (Path B (F x) (F y)) (mapOnPath A B F x y))\r
+ (\(x : A) -> propPi A (\(y : A) -> isEquiv (Path A x y) (Path B (F x) (F y)) (mapOnPath A B F x y))\r
+ (\(y : A) -> propIsEquiv (Path A x y) (Path B (F x) (F y)) (mapOnPath A B F x y)))\r
\r
propP2 (A B : U) (F : A -> B) : prop (P2 A B F)\r
- = propPi A (\(x : A) -> isContr ((y : A) * Id B (F x) (F y)))\r
- (\(x : A) -> propIsContr ((y : A) * Id B (F x) (F y)))\r
+ = propPi A (\(x : A) -> isContr ((y : A) * Path B (F x) (F y)))\r
+ (\(x : A) -> propIsContr ((y : A) * Path B (F x) (F y)))\r
\r
propP3 (A B : U) (F : A -> B) : prop (P3 A B F)\r
- = propPi B (\(x : B) -> prop ((y : A) * Id B x (F y)))\r
- (\(x : B) -> propIsProp ((y : A) * Id B x (F y)))\r
+ = propPi B (\(x : B) -> prop ((y : A) * Path B x (F y)))\r
+ (\(x : B) -> propIsProp ((y : A) * Path B x (F y)))\r
\r
-isoIdProp (A B : U) (AProp : prop A) (BProp : prop B) (F : A -> B) (G : B -> A) : Id U A B =\r
- isoId A B F G (\(y : B) -> BProp (F (G y)) y) (\(x : A) -> AProp (G (F x)) x)\r
+isoPathProp (A B : U) (AProp : prop A) (BProp : prop B) (F : A -> B) (G : B -> A) : Path U A B =\r
+ isoPath A B F G (\(y : B) -> BProp (F (G y)) y) (\(x : A) -> AProp (G (F x)) x)\r
\r
equivProp (A B : U) (AProp : prop A) (BProp : prop B) (F : A -> B) (G : B -> A) : equiv A B =\r
(F, gradLemma A B F G (\(y : B) -> BProp (F (G y)) y) (\(x : A) -> AProp (G (F x)) x))\r
\r
-isContrProp (A : U) (p : isContr A) (x y : A) : Id A x y = compId A x p.1 y (<i> p.2 x @ -i) (p.2 y)\r
+isContrProp (A : U) (p : isContr A) (x y : A) : Path A x y = compPath A x p.1 y (<i> p.2 x @ -i) (p.2 y)\r
\r
-lem2 (A : U) (x : A) : isContr ((y : A) * Id A x y)\r
+lem2 (A : U) (x : A) : isContr ((y : A) * Path A x y)\r
= ( (x, refl A x)\r
- , \(z : (y : A) * Id A x y) ->\r
- J A x (\(y : A) -> \(p : Id A x y) -> Id ((y : A) * Id A x y) (x, refl A x) (y, p))\r
- (refl ((y : A) * Id A x y) (x, refl A x)) z.1 z.2\r
+ , \(z : (y : A) * Path A x y) ->\r
+ J A x (\(y : A) -> \(p : Path A x y) -> Path ((y : A) * Path A x y) (x, refl A x) (y, p))\r
+ (refl ((y : A) * Path A x y) (x, refl A x)) z.1 z.2\r
)\r
\r
-lem31192 (A : U) (P : A -> U) (aC : isContr A) : Id U (Sigma A P) (P aC.1) =\r
- isoId (Sigma A P) (P aC.1) F G FG GF\r
+lem31192 (A : U) (P : A -> U) (aC : isContr A) : Path U (Sigma A P) (P aC.1) =\r
+ isoPath (Sigma A P) (P aC.1) F G FG GF\r
where\r
F (a : Sigma A P) : P aC.1 = transport (<i> P ((aC.2 a.1) @ -i)) a.2\r
G (a : P aC.1) : Sigma A P = (aC.1, a)\r
- FG (a : P aC.1) : Id (P aC.1) (transport (<i> P ((aC.2 aC.1) @ -i)) a) a = hole\r
+ FG (a : P aC.1) : Path (P aC.1) (transport (<i> P ((aC.2 aC.1) @ -i)) a) a = hole\r
where\r
- prf : Id (Id A aC.1 aC.1) (aC.2 aC.1) (<_> aC.1) = propSet A (isContrProp A aC) aC.1 aC.1 (aC.2 aC.1) (<_> aC.1)\r
- hole1 : Id (P aC.1) (transport (<_> P aC.1) a) a = transRefl (P aC.1) a\r
- hole : Id (P aC.1) (transport (<i> P ((aC.2 aC.1) @ -i)) a) a\r
- = transport (<i> Id (P aC.1) (transport (<j> P ((prf @ -i) @ -j)) a) a) hole1\r
- GF (a : Sigma A P) : Id (Sigma A P) (aC.1, transport (<i> P ((aC.2 a.1) @ -i)) a.2) a = hole\r
+ prf : Path (Path A aC.1 aC.1) (aC.2 aC.1) (<_> aC.1) = propSet A (isContrProp A aC) aC.1 aC.1 (aC.2 aC.1) (<_> aC.1)\r
+ hole1 : Path (P aC.1) (transport (<_> P aC.1) a) a = transRefl (P aC.1) a\r
+ hole : Path (P aC.1) (transport (<i> P ((aC.2 aC.1) @ -i)) a) a\r
+ = transport (<i> Path (P aC.1) (transport (<j> P ((prf @ -i) @ -j)) a) a) hole1\r
+ GF (a : Sigma A P) : Path (Sigma A P) (aC.1, transport (<i> P ((aC.2 a.1) @ -i)) a.2) a = hole\r
where\r
- hole2 : Id A aC.1 a.1 = aC.2 a.1\r
- hole1 : IdP (<i> P (hole2 @ i)) (transport (<i> P ((aC.2 a.1) @ -i)) a.2) a.2\r
+ hole2 : Path A aC.1 a.1 = aC.2 a.1\r
+ hole1 : PathP (<i> P (hole2 @ i)) (transport (<i> P ((aC.2 a.1) @ -i)) a.2) a.2\r
= <i> comp (<j> P (hole2 @ i \/ -j)) a.2 [(i=1) -> <_> a.2]\r
- hole : Id (Sigma A P) (aC.1, transport (<i> P ((aC.2 a.1) @ -i)) a.2) a\r
- = transport (<i> (lemIdSig A P (aC.1, transport (<i> P ((aC.2 a.1) @ -i)) a.2) a) @ -i) (hole2, hole1)\r
+ hole : Path (Sigma A P) (aC.1, transport (<i> P ((aC.2 a.1) @ -i)) a.2) a\r
+ = transport (<i> (lemPathSig A P (aC.1, transport (<i> P ((aC.2 a.1) @ -i)) a.2) a) @ -i) (hole2, hole1)\r
\r
total (A : U) (P Q : A -> U) (f : (x : A) -> P x -> Q x) (a : (x : A) * P x) : (x : A) * Q x = (a.1, f a.1 a.2)\r
\r
-ex210 (A : U) (B : A -> U) (C : (x : A) -> B x -> U) : Id U ((x : A) * (y : B x) * C x y) ((x : Sigma A B) * C x.1 x.2)\r
- = isoId ((x : A) * (y : B x) * C x y) ((x : Sigma A B) * C x.1 x.2) F G FG GF\r
+ex210 (A : U) (B : A -> U) (C : (x : A) -> B x -> U) : Path U ((x : A) * (y : B x) * C x y) ((x : Sigma A B) * C x.1 x.2)\r
+ = isoPath ((x : A) * (y : B x) * C x y) ((x : Sigma A B) * C x.1 x.2) F G FG GF\r
where\r
F (a : (x : A) * (y : B x) * C x y) : ((x : Sigma A B) * C x.1 x.2) = ((a.1, a.2.1), a.2.2)\r
G (a : (x : Sigma A B) * C x.1 x.2) : ((x : A) * (y : B x) * C x y) = (a.1.1, (a.1.2, a.2))\r
- FG (a : (x : Sigma A B) * C x.1 x.2) : Id ((x : Sigma A B) * C x.1 x.2) (F (G a)) a = <_> a\r
- GF (a : (x : A) * (y : B x) * C x y) : Id ((x : A) * (y : B x) * C x y) (G (F a)) a = <_> a\r
+ FG (a : (x : Sigma A B) * C x.1 x.2) : Path ((x : Sigma A B) * C x.1 x.2) (F (G a)) a = <_> a\r
+ GF (a : (x : A) * (y : B x) * C x y) : Path ((x : A) * (y : B x) * C x y) (G (F a)) a = <_> a\r
\r
-cSigma (A : U) (B : U) (C : A -> B -> U) : Id U ((x : A) * (y : B) * C x y) ((y : B) * (x : A) * C x y) =\r
- isoId ((x : A) * (y : B) * C x y) ((y : B) * (x : A) * C x y)\r
+cSigma (A : U) (B : U) (C : A -> B -> U) : Path U ((x : A) * (y : B) * C x y) ((y : B) * (x : A) * C x y) =\r
+ isoPath ((x : A) * (y : B) * C x y) ((y : B) * (x : A) * C x y)\r
(\(a : (x : A) * (y : B) * C x y) -> (a.2.1, (a.1, a.2.2)))\r
(\(a : (y : B) * (x : A) * C x y) -> (a.2.1, (a.1, a.2.2)))\r
(\(a : (y : B) * (x : A) * C x y) -> <_> a)\r
(\(a : (x : A) * (y : B) * C x y) -> <_> a)\r
\r
th476 (A : U) (P Q : A -> U) (f : (x : A) -> P x -> Q x) (x : A) (v : Q x)\r
- : Id U (fiber (Sigma A P) (Sigma A Q) (total A P Q f) (x, v)) (fiber (P x) (Q x) (f x) v)\r
+ : Path U (fiber (Sigma A P) (Sigma A Q) (total A P Q f) (x, v)) (fiber (P x) (Q x) (f x) v)\r
= hole\r
where\r
- A1 : U = (w : Sigma A P) * Id (Sigma A Q) (x, v) (total A P Q f w)\r
- A2 : U = (a : A) * (u : P a) * Id (Sigma A Q) (x, v) (a, f a u)\r
- A3 : U = (a : A) * (u : P a) * (p : Id A x a) * IdP (<i> Q (p @ i)) v (f a u)\r
- A4 : U = (a : A) * (p : Id A x a) * (u : P a) * IdP (<i> Q (p @ i)) v (f a u)\r
- A5 : U = (w : (a : A) * Id A x a) * (u : P w.1) * IdP (<i> Q (w.2 @ i)) v (f w.1 u)\r
- A6 : U = (u : P x) * Id (Q x) v (f x u)\r
- E12 : Id U A1 A2 = <i> (ex210 A P (\(a : A) -> \(b : P a) -> Id (Sigma A Q) (x, v) (a, f a b))) @ -i\r
- E23 : Id U A2 A3 = <i> (a : A) * (u : P a) * (lemIdSig A Q (x, v) (a, f a u)) @ i\r
- E34 : Id U A3 A4 = <i> (a : A) * (cSigma (P a) (Id A x a) (\(u : P a) -> \(p : Id A x a) -> IdP (<j> Q (p @ j)) v (f a u))) @ i\r
- E45 : Id U A4 A5 = ex210 A (Id A x) (\(a : A) -> \(p : Id A x a) -> (u : P a) * IdP (<i> Q (p @ i)) v (f a u))\r
- E56 : Id U A5 A6 = lem31192 ((a : A) * Id A x a) (\(w : (a : A) * Id A x a) -> (u : P w.1) * IdP (<i> Q (w.2 @ i)) v (f w.1 u))\r
+ A1 : U = (w : Sigma A P) * Path (Sigma A Q) (x, v) (total A P Q f w)\r
+ A2 : U = (a : A) * (u : P a) * Path (Sigma A Q) (x, v) (a, f a u)\r
+ A3 : U = (a : A) * (u : P a) * (p : Path A x a) * PathP (<i> Q (p @ i)) v (f a u)\r
+ A4 : U = (a : A) * (p : Path A x a) * (u : P a) * PathP (<i> Q (p @ i)) v (f a u)\r
+ A5 : U = (w : (a : A) * Path A x a) * (u : P w.1) * PathP (<i> Q (w.2 @ i)) v (f w.1 u)\r
+ A6 : U = (u : P x) * Path (Q x) v (f x u)\r
+ E12 : Path U A1 A2 = <i> (ex210 A P (\(a : A) -> \(b : P a) -> Path (Sigma A Q) (x, v) (a, f a b))) @ -i\r
+ E23 : Path U A2 A3 = <i> (a : A) * (u : P a) * (lemPathSig A Q (x, v) (a, f a u)) @ i\r
+ E34 : Path U A3 A4 = <i> (a : A) * (cSigma (P a) (Path A x a) (\(u : P a) -> \(p : Path A x a) -> PathP (<j> Q (p @ j)) v (f a u))) @ i\r
+ E45 : Path U A4 A5 = ex210 A (Path A x) (\(a : A) -> \(p : Path A x a) -> (u : P a) * PathP (<i> Q (p @ i)) v (f a u))\r
+ E56 : Path U A5 A6 = lem31192 ((a : A) * Path A x a) (\(w : (a : A) * Path A x a) -> (u : P w.1) * PathP (<i> Q (w.2 @ i)) v (f w.1 u))\r
(lem2 A x)\r
- hole : Id U A1 A6 = compId U A1 A2 A6 E12 (compId U A2 A3 A6 E23 (compId U A3 A4 A6 E34 (compId U A4 A5 A6 E45 E56)))\r
+ hole : Path U A1 A6 = compPath U A1 A2 A6 E12 (compPath U A2 A3 A6 E23 (compPath U A3 A4 A6 E34 (compPath U A4 A5 A6 E45 E56)))\r
\r
thm477 (A : U) (P Q : A -> U) (f : (x : A) -> P x -> Q x)\r
- : Id U ((x : A) -> isEquiv (P x) (Q x) (f x)) (isEquiv ((x : A) * P x) ((x : A) * Q x) (total A P Q f))\r
+ : Path U ((x : A) -> isEquiv (P x) (Q x) (f x)) (isEquiv ((x : A) * P x) ((x : A) * Q x) (total A P Q f))\r
= hole\r
where\r
AProp : prop ((x : A) -> isEquiv (P x) (Q x) (f x))\r
= transport (<i> isContr (th476 A P Q f y.1 y.2 @ -i)) (a y.1 y.2)\r
G (a : isEquiv ((x : A) * P x) ((x : A) * Q x) (total A P Q f)) (x : A) (y : Q x) : isContr (fiber (P x) (Q x) (f x) y)\r
= transport (<i> isContr (th476 A P Q f x y @ i)) (a (x, y))\r
- hole : Id U ((x : A) -> isEquiv (P x) (Q x) (f x)) (isEquiv ((x : A) * P x) ((x : A) * Q x) (total A P Q f))\r
- = isoIdProp ((x : A) -> isEquiv (P x) (Q x) (f x)) (isEquiv ((x : A) * P x) ((x : A) * Q x) (total A P Q f)) AProp BProp F G\r
+ hole : Path U ((x : A) -> isEquiv (P x) (Q x) (f x)) (isEquiv ((x : A) * P x) ((x : A) * Q x) (total A P Q f))\r
+ = isoPathProp ((x : A) -> isEquiv (P x) (Q x) (f x)) (isEquiv ((x : A) * P x) ((x : A) * Q x) (total A P Q f)) AProp BProp F G\r
\r
-F12 (A B : U) (F : A -> B) (p1 : P1 A B F) (x : A) : isContr ((y : A) * Id B (F x) (F y)) = hole\r
+F12 (A B : U) (F : A -> B) (p1 : P1 A B F) (x : A) : isContr ((y : A) * Path B (F x) (F y)) = hole\r
where\r
- hole3 : ((y : A) * Id A x y) -> ((y : A) * Id B (F x) (F y))\r
- = total A (\(y : A) -> Id A x y) (\(y : A) -> Id B (F x) (F y)) (mapOnPath A B F x)\r
- hole2 : isEquiv ((y : A) * Id A x y) ((y : A) * Id B (F x) (F y)) hole3\r
- = transport (thm477 A (\(y : A) -> Id A x y) (\(y : A) -> Id B (F x) (F y)) (mapOnPath A B F x)) (p1 x)\r
- hole4 : Id U ((y : A) * Id A x y) ((y : A) * Id B (F x) (F y)) = equivId ((y : A) * Id A x y) ((y : A) * Id B (F x) (F y)) hole3 hole2\r
- hole : isContr ((y : A) * Id B (F x) (F y)) = transport (<i> isContr (hole4@i)) (lem2 A x)\r
+ hole3 : ((y : A) * Path A x y) -> ((y : A) * Path B (F x) (F y))\r
+ = total A (\(y : A) -> Path A x y) (\(y : A) -> Path B (F x) (F y)) (mapOnPath A B F x)\r
+ hole2 : isEquiv ((y : A) * Path A x y) ((y : A) * Path B (F x) (F y)) hole3\r
+ = transport (thm477 A (\(y : A) -> Path A x y) (\(y : A) -> Path B (F x) (F y)) (mapOnPath A B F x)) (p1 x)\r
+ hole4 : Path U ((y : A) * Path A x y) ((y : A) * Path B (F x) (F y)) = equivPath ((y : A) * Path A x y) ((y : A) * Path B (F x) (F y)) hole3 hole2\r
+ hole : isContr ((y : A) * Path B (F x) (F y)) = transport (<i> isContr (hole4@i)) (lem2 A x)\r
\r
contrEquiv (A B : U) (aC : isContr A) (bC : isContr B) : equiv A B\r
= (\(_ : A) -> bC.1, gradLemma A B (\(_ : A) -> bC.1) (\(_ : B) -> aC.1) (\(x : B) -> bC.2 x) (\(x : A) -> aC.2 x))\r
\r
-F21 (A B : U) (F : A -> B) (p2 : P2 A B F) (x y : A) : isEquiv (Id A x y) (Id B (F x) (F y)) (mapOnPath A B F x y) = hole3 y\r
+F21 (A B : U) (F : A -> B) (p2 : P2 A B F) (x y : A) : isEquiv (Path A x y) (Path B (F x) (F y)) (mapOnPath A B F x y) = hole3 y\r
where\r
- hole0 : ((y : A) * Id A x y) -> ((y : A) * Id B (F x) (F y))\r
- = total A (\(y : A) -> Id A x y) (\(y : A) -> Id B (F x) (F y)) (mapOnPath A B F x)\r
- hole2 : isEquiv ((y : A) * Id A x y) ((y : A) * Id B (F x) (F y)) hole0\r
- = (equivProp ((y : A) * Id A x y) ((y : A) * Id B (F x) (F y))\r
- (isContrProp ((y : A) * Id A x y) (lem2 A x))\r
- (isContrProp ((y : A) * Id B (F x) (F y)) (p2 x))\r
- hole0 (\(_ : (y : A) * Id B (F x) (F y)) -> (x, <_> x))).2\r
- hole4 : Id U ((y : A) -> isEquiv (Id A x y) (Id B (F x) (F y)) (mapOnPath A B F x y)) (isEquiv ((y : A) * Id A x y) ((y : A) * Id B (F x) (F y)) hole0)\r
- = thm477 A (\(y : A) -> Id A x y) (\(y : A) -> Id B (F x) (F y)) (mapOnPath A B F x)\r
- hole3 : (y : A) -> isEquiv (Id A x y) (Id B (F x) (F y)) (mapOnPath A B F x y)\r
+ hole0 : ((y : A) * Path A x y) -> ((y : A) * Path B (F x) (F y))\r
+ = total A (\(y : A) -> Path A x y) (\(y : A) -> Path B (F x) (F y)) (mapOnPath A B F x)\r
+ hole2 : isEquiv ((y : A) * Path A x y) ((y : A) * Path B (F x) (F y)) hole0\r
+ = (equivProp ((y : A) * Path A x y) ((y : A) * Path B (F x) (F y))\r
+ (isContrProp ((y : A) * Path A x y) (lem2 A x))\r
+ (isContrProp ((y : A) * Path B (F x) (F y)) (p2 x))\r
+ hole0 (\(_ : (y : A) * Path B (F x) (F y)) -> (x, <_> x))).2\r
+ hole4 : Path U ((y : A) -> isEquiv (Path A x y) (Path B (F x) (F y)) (mapOnPath A B F x y)) (isEquiv ((y : A) * Path A x y) ((y : A) * Path B (F x) (F y)) hole0)\r
+ = thm477 A (\(y : A) -> Path A x y) (\(y : A) -> Path B (F x) (F y)) (mapOnPath A B F x)\r
+ hole3 : (y : A) -> isEquiv (Path A x y) (Path B (F x) (F y)) (mapOnPath A B F x y)\r
= transport (<i> hole4 @ -i) hole2\r
\r
-F32 (A B : U) (F : A -> B) (p3 : P3 A B F) (x : A) : isContr ((y : A) * Id B (F x) (F y))\r
- = ((x, refl B (F x)), \(q : ((y : A) * Id B (F x) (F y))) -> p3 (F x) (x, refl B (F x)) q)\r
+F32 (A B : U) (F : A -> B) (p3 : P3 A B F) (x : A) : isContr ((y : A) * Path B (F x) (F y))\r
+ = ((x, refl B (F x)), \(q : ((y : A) * Path B (F x) (F y))) -> p3 (F x) (x, refl B (F x)) q)\r
\r
-F23 (A B : U) (F : A -> B) (p2 : P2 A B F) (x : B) (a b : (y : A) * Id B x (F y)) : Id ((y : A) * Id B x (F y)) a b = hole\r
+F23 (A B : U) (F : A -> B) (p2 : P2 A B F) (x : B) (a b : (y : A) * Path B x (F y)) : Path ((y : A) * Path B x (F y)) a b = hole\r
where\r
- hole2 : Id ((y : A) * Id B (F a.1) (F y)) (a.1, refl B (F a.1)) (b.1, compId B (F a.1) x (F b.1) (<i> a.2 @ -i) b.2)\r
- = isContrProp ((y : A) * Id B (F a.1) (F y)) (p2 (a.1)) (a.1, refl B (F a.1)) (b.1, compId B (F a.1) x (F b.1) (<i> a.2 @ -i) b.2)\r
- hole3 : (Id ((y : A) * Id B x (F y)) a (b.1, compId B x x (F b.1) (<_> x) b.2))\r
+ hole2 : Path ((y : A) * Path B (F a.1) (F y)) (a.1, refl B (F a.1)) (b.1, compPath B (F a.1) x (F b.1) (<i> a.2 @ -i) b.2)\r
+ = isContrProp ((y : A) * Path B (F a.1) (F y)) (p2 (a.1)) (a.1, refl B (F a.1)) (b.1, compPath B (F a.1) x (F b.1) (<i> a.2 @ -i) b.2)\r
+ hole3 : (Path ((y : A) * Path B x (F y)) a (b.1, compPath B x x (F b.1) (<_> x) b.2))\r
= transport\r
- (<i> Id ((y : A) * Id B (a.2 @ -i) (F y)) (a.1, <j> a.2 @ -i \/ j) (b.1, compId B (a.2 @ -i) x (F b.1) (<j> a.2 @ -i /\ -j) b.2))\r
+ (<i> Path ((y : A) * Path B (a.2 @ -i) (F y)) (a.1, <j> a.2 @ -i \/ j) (b.1, compPath B (a.2 @ -i) x (F b.1) (<j> a.2 @ -i /\ -j) b.2))\r
hole2\r
- hole : Id ((y : A) * Id B x (F y)) a b\r
+ hole : Path ((y : A) * Path B x (F y)) a b\r
= transport\r
- (<i> Id ((y : A) * Id B x (F y)) a (b.1, (lemReflComp B x (F b.1) b.2) @ i))\r
+ (<i> Path ((y : A) * Path B x (F y)) a (b.1, (lemReflComp B x (F b.1) b.2) @ i))\r
hole3\r
\r
-E12 (A B : U) (F : A -> B) : Id U (P1 A B F) (P2 A B F) = isoIdProp (P1 A B F) (P2 A B F) (propP1 A B F) (propP2 A B F) (F12 A B F) (F21 A B F)\r
+E12 (A B : U) (F : A -> B) : Path U (P1 A B F) (P2 A B F) = isoPathProp (P1 A B F) (P2 A B F) (propP1 A B F) (propP2 A B F) (F12 A B F) (F21 A B F)\r
opaque E12\r
-E23 (A B : U) (F : A -> B) : Id U (P2 A B F) (P3 A B F) = isoIdProp (P2 A B F) (P3 A B F) (propP2 A B F) (propP3 A B F) (F23 A B F) (F32 A B F)\r
+E23 (A B : U) (F : A -> B) : Path U (P2 A B F) (P3 A B F) = isoPathProp (P2 A B F) (P3 A B F) (propP2 A B F) (propP3 A B F) (F23 A B F) (F32 A B F)\r
opaque E23\r
-E13 (A B : U) (F : A -> B) : Id U (P1 A B F) (P3 A B F) = compId U (P1 A B F) (P2 A B F) (P3 A B F) (E12 A B F) (E23 A B F)\r
+E13 (A B : U) (F : A -> B) : Path U (P1 A B F) (P3 A B F) = compPath U (P1 A B F) (P2 A B F) (P3 A B F) (E12 A B F) (E23 A B F)\r
opaque E13\r
\r
-- torsor\r
inl n -> hole n\r
where hole : (n : nat) -> isEquiv A A (\(x : A) -> action A shift x (inl n)) = split\r
zero -> hole0\r
- where hole0 : isEquiv A A (\(x : A) -> (shift.2 x).1.1) = isEquivComp' A A A (\(x : A) -> x) shift.1 (isEquivId A) shift.2\r
+ where hole0 : isEquiv A A (\(x : A) -> (shift.2 x).1.1) = isEquivComp' A A A (\(x : A) -> x) shift.1 (isEquivPath A) shift.2\r
suc n -> isEquivComp' A A A (\(x : A) -> action A shift x (inl n)) shift.1 (hole n) shift.2\r
inr n -> hole n\r
where hole : (n : nat) -> isEquiv A A (\(x : A) -> action A shift x (inr n)) = split\r
- zero -> isEquivId A\r
+ zero -> isEquivPath A\r
suc n -> isEquivComp A A A (\(x : A) -> action A shift x (inr n)) shift.1 (hole n) shift.2\r
\r
BZ : U = (A : U) * (a : ishinh_UU A)\r
BZP (A : BZ) (a : BZSet A) : BZSet A = ((BZShift A).2 a).1.1\r
BZEquiv (A : BZ) : (x : BZSet A) -> isEquiv Z (BZSet A) (BZAction A x) = A.2.2.2\r
BZASet (x : BZ) : set (BZSet x) = BZNE x (set (BZSet x), setIsProp (BZSet x))\r
- (\(a : BZSet x) -> transport (<i> set (equivId Z (BZSet x) (BZAction x a) (BZEquiv x a) @ i)) ZSet)\r
+ (\(a : BZSet x) -> transport (<i> set (equivPath Z (BZSet x) (BZAction x a) (BZEquiv x a) @ i)) ZSet)\r
\r
-- Two Z-torsors are equal if their underlying sets are equal in a way compatible with the actions\r
-lemBZ (BA BB : BZ) : equiv ((p : Id U (BZSet BA) (BZSet BB)) * IdP (<i> p@i -> p@i) (BZS BA) (BZS BB)) (Id BZ BA BB) = hole\r
+lemBZ (BA BB : BZ) : equiv ((p : Path U (BZSet BA) (BZSet BB)) * PathP (<i> p@i -> p@i) (BZS BA) (BZS BB)) (Path BZ BA BB) = hole\r
where\r
A : U = BA.1\r
a : ishinh_UU A = BA.2.1\r
b : ishinh_UU B = BB.2.1\r
BShift : equiv B B = BB.2.2.1\r
BEquiv : (x : B) -> isEquiv Z B (BZAction BB x) = BB.2.2.2\r
- F (w : (p : Id U A B) * IdP (<i> p@i -> p@i) (BZS BA) (BZS BB)) : Id BZ BA BB = hole\r
+ F (w : (p : Path U A B) * PathP (<i> p@i -> p@i) (BZS BA) (BZS BB)) : Path BZ BA BB = hole\r
where\r
- p : Id U A B = w.1\r
- pS : IdP (<i> p@i -> p@i) (BZS BA) (BZS BB) = w.2\r
- pNE : IdP (<i> ishinh_UU (p@i)) a b\r
- = lemIdPProp (ishinh_UU A) (ishinh_UU B) (propishinh A) (<i> ishinh_UU (p@i)) a b\r
- pShift : IdP (<i> equiv (p@i) (p@i)) AShift BShift =\r
+ p : Path U A B = w.1\r
+ pS : PathP (<i> p@i -> p@i) (BZS BA) (BZS BB) = w.2\r
+ pNE : PathP (<i> ishinh_UU (p@i)) a b\r
+ = lemPathPProp (ishinh_UU A) (ishinh_UU B) (propishinh A) (<i> ishinh_UU (p@i)) a b\r
+ pShift : PathP (<i> equiv (p@i) (p@i)) AShift BShift =\r
<i> ( pS @ i\r
- , (lemIdPProp\r
+ , (lemPathPProp\r
(isEquiv A A AShift.1)\r
(isEquiv B B BShift.1)\r
(propIsEquiv A A AShift.1)\r
(<j> isEquiv (p@j) (p@j) (pS@j))\r
AShift.2 BShift.2) @ i\r
)\r
- pEquiv : IdP (<i> (x : p@i) -> isEquiv Z (p@i) (action (p@i) (pShift@i) x)) AEquiv BEquiv\r
- = lemIdPProp\r
+ pEquiv : PathP (<i> (x : p@i) -> isEquiv Z (p@i) (action (p@i) (pShift@i) x)) AEquiv BEquiv\r
+ = lemPathPProp\r
((x : A) -> isEquiv Z A (action A AShift x))\r
((x : B) -> isEquiv Z B (action B BShift x))\r
(propPi A (\(x : A) -> isEquiv Z A (action A AShift x))\r
(\(x : A) -> propIsEquiv Z A (action A AShift x)))\r
(<i> (x : p@i) -> isEquiv Z (p@i) (action (p@i) (pShift@i) x))\r
AEquiv BEquiv\r
- hole : Id BZ BA BB = <i> (p@i, (pNE@i, (pShift@i, pEquiv@i)))\r
- G (q : Id BZ BA BB) : (p : Id U A B) * IdP (<i> p@i -> p@i) (BZS BA) (BZS BB) = (<i> BZSet (q @ i), <i> (BZShift (q@i)).1)\r
- GF (w : (p : Id U A B) * IdP (<i> p@i -> p@i) (BZS BA) (BZS BB))\r
- : Id ((p : Id U A B) * IdP (<i> p@i -> p@i) (BZS BA) (BZS BB)) (<i> BZSet (F w @ i), <i> (BZShift (F w @ i)).1) w\r
+ hole : Path BZ BA BB = <i> (p@i, (pNE@i, (pShift@i, pEquiv@i)))\r
+ G (q : Path BZ BA BB) : (p : Path U A B) * PathP (<i> p@i -> p@i) (BZS BA) (BZS BB) = (<i> BZSet (q @ i), <i> (BZShift (q@i)).1)\r
+ GF (w : (p : Path U A B) * PathP (<i> p@i -> p@i) (BZS BA) (BZS BB))\r
+ : Path ((p : Path U A B) * PathP (<i> p@i -> p@i) (BZS BA) (BZS BB)) (<i> BZSet (F w @ i), <i> (BZShift (F w @ i)).1) w\r
= <_> w\r
- FG (q : Id BZ BA BB) : Id (Id BZ BA BB) (F (<i> BZSet (q@i), <i> (BZShift (q@i)).1)) q = hole\r
+ FG (q : Path BZ BA BB) : Path (Path BZ BA BB) (F (<i> BZSet (q@i), <i> (BZShift (q@i)).1)) q = hole\r
where\r
- p : Id U A B = <i> BZSet (q@i)\r
- p2 : IdP (<i> p@i -> p@i) (BZS BA) (BZS BB) = <i> (BZShift (q@i)).1\r
- q2 : Id BZ BA BB = F (p, p2)\r
- pNE : Id (IdP (<i> ishinh_UU (p@i)) a b) (<i> BZNE (q2@i)) (<i> BZNE (q@i))\r
- = lemIdPSet (ishinh_UU A) (ishinh_UU B) (propSet (ishinh_UU A) (propishinh A)) (<i> ishinh_UU (p@i)) a b (<i> BZNE (q2@i)) (<i> BZNE (q@i))\r
- pShift : Id (IdP (<i> equiv (p@i) (p@i)) AShift BShift) (<i> BZShift (q2@i)) (<i> BZShift (q@i)) =\r
+ p : Path U A B = <i> BZSet (q@i)\r
+ p2 : PathP (<i> p@i -> p@i) (BZS BA) (BZS BB) = <i> (BZShift (q@i)).1\r
+ q2 : Path BZ BA BB = F (p, p2)\r
+ pNE : Path (PathP (<i> ishinh_UU (p@i)) a b) (<i> BZNE (q2@i)) (<i> BZNE (q@i))\r
+ = lemPathPSet (ishinh_UU A) (ishinh_UU B) (propSet (ishinh_UU A) (propishinh A)) (<i> ishinh_UU (p@i)) a b (<i> BZNE (q2@i)) (<i> BZNE (q@i))\r
+ pShift : Path (PathP (<i> equiv (p@i) (p@i)) AShift BShift) (<i> BZShift (q2@i)) (<i> BZShift (q@i)) =\r
<j i> ( p2 @ i\r
- , (lemIdPSet\r
+ , (lemPathPSet\r
(isEquiv A A AShift.1)\r
(isEquiv B B BShift.1)\r
(propSet (isEquiv A A AShift.1) (propIsEquiv A A AShift.1))\r
AShift.2 BShift.2\r
(<i> (BZShift (q2@i)).2) (<i> (BZShift (q@i)).2)) @ j @ i\r
)\r
- pEquiv : IdP (<j> IdP (<i> (x : p@i) -> isEquiv Z (p@i) (action (p@i) (pShift@j@i) x)) AEquiv BEquiv) (<i> BZEquiv (q2@i)) (<i> BZEquiv (q@i))\r
- = lemIdPSet2\r
+ pEquiv : PathP (<j> PathP (<i> (x : p@i) -> isEquiv Z (p@i) (action (p@i) (pShift@j@i) x)) AEquiv BEquiv) (<i> BZEquiv (q2@i)) (<i> BZEquiv (q@i))\r
+ = lemPathPSet2\r
((x : A) -> isEquiv Z A (action A AShift x))\r
((x : B) -> isEquiv Z B (action B BShift x))\r
(propSet ((x : A) -> isEquiv Z A (action A AShift x))\r
(<i> (x : p@i) -> isEquiv Z (p@i) (action (p@i) (pShift@1@i) x))\r
(<j i> (x : p@i) -> isEquiv Z (p@i) (action (p@i) (pShift@j@i) x))\r
AEquiv BEquiv (<i> BZEquiv (q2@i)) (<i> BZEquiv (q@i))\r
- hole : Id (Id BZ BA BB) q2 q = <j i> (p@i, (pNE@j@i, (pShift@j@i, pEquiv@j@i)))\r
- hole : equiv ((p : Id U (BZSet BA) (BZSet BB)) * IdP (<i> p@i -> p@i) (BZS BA) (BZS BB)) (Id BZ BA BB)\r
- = (F, gradLemma ((p : Id U (BZSet BA) (BZSet BB)) * IdP (<i> p@i -> p@i) (BZS BA) (BZS BB)) (Id BZ BA BB) F G FG GF)\r
+ hole : Path (Path BZ BA BB) q2 q = <j i> (p@i, (pNE@j@i, (pShift@j@i, pEquiv@j@i)))\r
+ hole : equiv ((p : Path U (BZSet BA) (BZSet BB)) * PathP (<i> p@i -> p@i) (BZS BA) (BZS BB)) (Path BZ BA BB)\r
+ = (F, gradLemma ((p : Path U (BZSet BA) (BZSet BB)) * PathP (<i> p@i -> p@i) (BZS BA) (BZS BB)) (Path BZ BA BB) F G FG GF)\r
\r
-lem2 (A : U) (shift : equiv A A) (a : A) : (x : Z) -> Id A (shift.1 (action A shift a x)) (action A shift a (sucZ x)) = split\r
+lem2 (A : U) (shift : equiv A A) (a : A) : (x : Z) -> Path A (shift.1 (action A shift a x)) (action A shift a (sucZ x)) = split\r
inl n -> lem2Aux n\r
where\r
- lem2Aux : (n : nat) -> Id A (shift.1 (action A shift a (inl n))) (action A shift a (sucZ (inl n))) = split\r
+ lem2Aux : (n : nat) -> Path A (shift.1 (action A shift a (inl n))) (action A shift a (sucZ (inl n))) = split\r
zero -> retEq A A shift a\r
suc n -> retEq A A shift (action A shift a (inl n))\r
inr n -> lem2Aux n\r
where\r
- lem2Aux : (n : nat) -> Id A (shift.1 (action A shift a (inr n))) (action A shift a (sucZ (inr n))) = split\r
+ lem2Aux : (n : nat) -> Path A (shift.1 (action A shift a (inr n))) (action A shift a (sucZ (inr n))) = split\r
zero -> <_> shift.1 a\r
suc n -> <_> shift.1 (action A shift a (inr (suc n)))\r
\r
-lem2' (A : U) (shift : equiv A A) (a : A) : (x : Z) -> Id A (shift.2 (action A shift a x)).1.1 (action A shift a (predZ x)) = split\r
+lem2' (A : U) (shift : equiv A A) (a : A) : (x : Z) -> Path A (shift.2 (action A shift a x)).1.1 (action A shift a (predZ x)) = split\r
inl n -> lem2Aux n\r
where\r
- lem2Aux : (n : nat) -> Id A (shift.2 (action A shift a (inl n))).1.1 (action A shift a (predZ (inl n))) = split\r
+ lem2Aux : (n : nat) -> Path A (shift.2 (action A shift a (inl n))).1.1 (action A shift a (predZ (inl n))) = split\r
zero -> <_> action A shift a (predZ (inl zero))\r
suc n -> <_> action A shift a (predZ (inl (suc n)))\r
inr n -> lem2Aux n\r
where\r
- lem2Aux : (n : nat) -> Id A (shift.2 (action A shift a (inr n))).1.1 (action A shift a (predZ (inr n))) = split\r
+ lem2Aux : (n : nat) -> Path A (shift.2 (action A shift a (inr n))).1.1 (action A shift a (predZ (inr n))) = split\r
zero -> <_> action A shift a (inl zero)\r
suc n -> secEq A A shift (action A shift a (inr n))\r
\r
where\r
ZShift : equiv Z Z = (sucZ, gradLemma Z Z sucZ predZ sucpredZ predsucZ)\r
plus : Z -> Z -> Z = action Z ZShift\r
- plusCommZero : (y : Z) -> Id Z (plus zeroZ y) (plus y zeroZ) = split\r
+ plusCommZero : (y : Z) -> Path Z (plus zeroZ y) (plus y zeroZ) = split\r
inl u -> hole u\r
- where hole : (n : nat) -> Id Z (plus zeroZ (inl n)) (inl n) = split\r
+ where hole : (n : nat) -> Path Z (plus zeroZ (inl n)) (inl n) = split\r
zero -> <_> inl zero\r
suc n -> <i> predZ (hole n @ i)\r
inr u -> hole u\r
- where hole : (n : nat) -> Id Z (plus zeroZ (inr n)) (inr n) = split\r
+ where hole : (n : nat) -> Path Z (plus zeroZ (inr n)) (inr n) = split\r
zero -> <_> inr zero\r
suc n -> <i> sucZ (plusCommZero (inr n) @ i)\r
- plusCommSuc (x : Z) : (y : Z) -> Id Z (plus (sucZ x) y) (plus x (sucZ y)) = split\r
+ plusCommSuc (x : Z) : (y : Z) -> Path Z (plus (sucZ x) y) (plus x (sucZ y)) = split\r
inl u -> hole u\r
- where hole : (n : nat) -> Id Z (plus (sucZ x) (inl n)) (plus x (sucZ (inl n))) = split\r
+ where hole : (n : nat) -> Path Z (plus (sucZ x) (inl n)) (plus x (sucZ (inl n))) = split\r
zero -> predsucZ x\r
suc n -> hole0 n\r
- where hole0 : (n : nat) -> Id Z (predZ (plus (sucZ x) (inl n))) (plus x (sucZ (predZ (inl n)))) = split\r
+ where hole0 : (n : nat) -> Path Z (predZ (plus (sucZ x) (inl n))) (plus x (sucZ (predZ (inl n)))) = split\r
zero -> <i> predZ (predsucZ x @ i)\r
- suc n -> compId Z (predZ (predZ (plus (sucZ x) (inl n))))\r
+ suc n -> compPath Z (predZ (predZ (plus (sucZ x) (inl n))))\r
(predZ (plus x (inl n))) (predZ (plus x (sucZ (predZ (inl n)))))\r
(<i> predZ (hole0 n @ i)) (<i> predZ (plus x (sucpredZ (inl n) @ -i)))\r
inr u -> hole u\r
- where hole : (n : nat) -> Id Z (plus (sucZ x) (inr n)) (plus x (inr (suc n))) = split\r
+ where hole : (n : nat) -> Path Z (plus (sucZ x) (inr n)) (plus x (inr (suc n))) = split\r
zero -> <_> sucZ x\r
suc n -> <i> sucZ (hole n @ i)\r
- plusCommPred (x y : Z) : Id Z (plus (predZ x) y) (plus x (predZ y))\r
- = transport (<i> Id Z (plus (predZ x) (sucpredZ y @ i)) (plus (sucpredZ x @ i) (predZ y)))\r
+ plusCommPred (x y : Z) : Path Z (plus (predZ x) y) (plus x (predZ y))\r
+ = transport (<i> Path Z (plus (predZ x) (sucpredZ y @ i)) (plus (sucpredZ x @ i) (predZ y)))\r
(<i> plusCommSuc (predZ x) (predZ y) @ -i)\r
- plusComm (x : Z) : (y : Z) -> Id Z (plus y x) (plus x y) = split\r
+ plusComm (x : Z) : (y : Z) -> Path Z (plus y x) (plus x y) = split\r
inl u -> hole u\r
- where hole : (n : nat) -> Id Z (plus (inl n) x) (plus x (inl n)) = split\r
- zero -> compId Z (plus (inl zero) x) (plus (inr zero) (predZ x)) (predZ x)\r
+ where hole : (n : nat) -> Path Z (plus (inl n) x) (plus x (inl n)) = split\r
+ zero -> compPath Z (plus (inl zero) x) (plus (inr zero) (predZ x)) (predZ x)\r
(plusCommPred (inr zero) x) (plusCommZero (predZ x))\r
- suc n -> compId Z (plus (inl (suc n)) x) (plus (inl n) (predZ x)) (predZ (plus x (inl n)))\r
+ suc n -> compPath Z (plus (inl (suc n)) x) (plus (inl n) (predZ x)) (predZ (plus x (inl n)))\r
(plusCommPred (inl n) x)\r
- (compId Z (plus (inl n) (predZ x)) (predZ (plus (inl n) x)) (predZ (plus x (inl n)))\r
+ (compPath Z (plus (inl n) (predZ x)) (predZ (plus (inl n) x)) (predZ (plus x (inl n)))\r
(<i> lem2' Z ZShift (inl n) x @ -i)\r
(<i> predZ (hole n @ i)))\r
inr u -> hole u\r
- where hole : (n : nat) -> Id Z (plus (inr n) x) (plus x (inr n)) = split\r
+ where hole : (n : nat) -> Path Z (plus (inr n) x) (plus x (inr n)) = split\r
zero -> plusCommZero x\r
- suc n -> compId Z (plus (inr (suc n)) x) (plus (inr n) (sucZ x)) (sucZ (plus x (inr n)))\r
+ suc n -> compPath Z (plus (inr (suc n)) x) (plus (inr n) (sucZ x)) (sucZ (plus x (inr n)))\r
(plusCommSuc (inr n) x)\r
- (compId Z (plus (inr n) (sucZ x)) (sucZ (plus (inr n) x)) (sucZ (plus x (inr n)))\r
+ (compPath Z (plus (inr n) (sucZ x)) (sucZ (plus (inr n) x)) (sucZ (plus x (inr n)))\r
(<i> lem2 Z ZShift (inr n) x @ -i)\r
(<i> sucZ (hole n @ i)))\r
ZEquiv (x : Z) : isEquiv Z Z (plus x)\r
= transport (<i> isEquiv Z Z (\(y : Z) -> plusComm x y @ i)) (actionEquiv Z ZShift x)\r
\r
-loopBZ : U = Id BZ ZBZ ZBZ\r
-compBZ : loopBZ -> loopBZ -> loopBZ = compId BZ ZBZ ZBZ ZBZ\r
+loopBZ : U = Path BZ ZBZ ZBZ\r
+compBZ : loopBZ -> loopBZ -> loopBZ = compPath BZ ZBZ ZBZ ZBZ\r
\r
-transportIsoId (A B : U) (f : A -> B) (g : B -> A)\r
- (s : (y : B) -> Id B (f (g y)) y)\r
- (t : (x : A) -> Id A (g (f x)) x)\r
+transportIsoPath (A B : U) (f : A -> B) (g : B -> A)\r
+ (s : (y : B) -> Path B (f (g y)) y)\r
+ (t : (x : A) -> Path A (g (f x)) x)\r
(x : A)\r
- : Id B (transport (isoId A B f g s t) x) (f x)\r
- = compId B (transport (<_> B) (transport (<_> B) (f x))) (transport (<_> B) (f x)) (f x)\r
+ : Path B (transport (isoPath A B f g s t) x) (f x)\r
+ = compPath B (transport (<_> B) (transport (<_> B) (f x))) (transport (<_> B) (f x)) (f x)\r
(<i> transport (<_> B) (transRefl B (f x) @ i)) (transRefl B (f x))\r
\r
-transportIsoId' (A B : U) (f : A -> B) (g : B -> A)\r
- (s : (y : B) -> Id B (f (g y)) y)\r
- (t : (x : A) -> Id A (g (f x)) x)\r
+transportIsoPath' (A B : U) (f : A -> B) (g : B -> A)\r
+ (s : (y : B) -> Path B (f (g y)) y)\r
+ (t : (x : A) -> Path A (g (f x)) x)\r
(x : B)\r
- : Id A (transport (<i> isoId A B f g s t @ -i) x) (g x)\r
- = compId A (transport (<_> A) (g (transport (<_> B) x))) (transport (<_> A) (g x)) (g x)\r
+ : Path A (transport (<i> isoPath A B f g s t @ -i) x) (g x)\r
+ = compPath A (transport (<_> A) (g (transport (<_> B) x))) (transport (<_> A) (g x)) (g x)\r
(<i> transport (<_> A) (g (transRefl B x @ i))) (transRefl A (g x))\r
\r
-loopA (A : BZ) : Id BZ A A = (lemBZ A A).1 (shiftId, hole)\r
+loopA (A : BZ) : Path BZ A A = (lemBZ A A).1 (shiftPath, hole)\r
where\r
AS : U = BZSet A\r
- shiftId : Id U AS AS = equivId AS AS (BZShift A).1 (BZShift A).2\r
- hole1 (x : AS) : Id AS (transport (<_>AS) (transport (<_> AS) (BZS A (BZS A (transport (<_> AS) (BZP A (transport (<_>AS) x))))))) (BZS A (BZS A (BZP A x)))\r
- = compId AS (transport (<_>AS) (transport (<_> AS) (BZS A (BZS A (transport (<_> AS) (BZP A (transport (<_>AS) x)))))))\r
+ shiftPath : Path U AS AS = equivPath AS AS (BZShift A).1 (BZShift A).2\r
+ hole1 (x : AS) : Path AS (transport (<_>AS) (transport (<_> AS) (BZS A (BZS A (transport (<_> AS) (BZP A (transport (<_>AS) x))))))) (BZS A (BZS A (BZP A x)))\r
+ = compPath AS (transport (<_>AS) (transport (<_> AS) (BZS A (BZS A (transport (<_> AS) (BZP A (transport (<_>AS) x)))))))\r
(BZS A (BZS A (transport (<_> AS) (BZP A (transport (<_>AS) x)))))\r
(BZS A (BZS A (BZP A x)))\r
(transConstNRefl AS (BZS A (BZS A (transport (<_> AS) (BZP A (transport (<_>AS) x))))) (suc (suc zero)))\r
- (compId AS (BZS A (BZS A (transport (<_> AS) (BZP A (transport (<_>AS) x)))))\r
+ (compPath AS (BZS A (BZS A (transport (<_> AS) (BZP A (transport (<_>AS) x)))))\r
(BZS A (BZS A (BZP A (transport (<_>AS) x))))\r
(BZS A (BZS A (BZP A x)))\r
(<i> BZS A (BZS A (transRefl AS (BZP A (transport (<_>AS) x)) @ i)))\r
(<i> BZS A (BZS A (BZP A (transRefl AS x @ i))))\r
)\r
- hole2 (x : AS) : Id AS (BZS A (BZS A (BZP A x))) (BZS A x) = <i> BZS A (retEq AS AS (BZShift A) x @ i)\r
- hole3 : Id (AS -> AS) (transport (<i> (shiftId@i) -> (shiftId@i)) (BZS A)) (BZS A)\r
- = <i> \(x : AS) -> compId AS (transport shiftId (BZS A (transport (<j> shiftId @ -j) x))) (BZS A (BZS A (BZP A x))) (BZS A x)\r
+ hole2 (x : AS) : Path AS (BZS A (BZS A (BZP A x))) (BZS A x) = <i> BZS A (retEq AS AS (BZShift A) x @ i)\r
+ hole3 : Path (AS -> AS) (transport (<i> (shiftPath@i) -> (shiftPath@i)) (BZS A)) (BZS A)\r
+ = <i> \(x : AS) -> compPath AS (transport shiftPath (BZS A (transport (<j> shiftPath @ -j) x))) (BZS A (BZS A (BZP A x))) (BZS A x)\r
(hole1 x) (hole2 x) @ i\r
- hole : IdP (<i> (shiftId@i) -> (shiftId@i)) (BZS A) (BZS A)\r
- = substIdP (AS->AS) (AS->AS) (<i> (shiftId@i) -> (shiftId@i)) (BZS A) (BZS A) hole3\r
+ hole : PathP (<i> (shiftPath@i) -> (shiftPath@i)) (BZS A) (BZS A)\r
+ = substPathP (AS->AS) (AS->AS) (<i> (shiftPath@i) -> (shiftPath@i)) (BZS A) (BZS A) hole3\r
\r
loopZ : loopBZ = loopA ZBZ\r
invLoopZ : loopBZ = <i> loopZ @ -i\r
\r
-- loopBZ = Z = loopS1\r
\r
-encodeZ (A : BZ) (p : Id BZ ZBZ A) : BZSet A = transport (<i> BZSet (p@i)) zeroZ\r
+encodeZ (A : BZ) (p : Path BZ ZBZ A) : BZSet A = transport (<i> BZSet (p@i)) zeroZ\r
\r
-decodeZP (A : BZ) (a : BZSet A) : Id U (BZSet A) Z = <i> (univalence (BZSet A) Z (BZAction A a, BZEquiv A a)).1.1 @ -i\r
+decodeZP (A : BZ) (a : BZSet A) : Path U (BZSet A) Z = <i> (univalence (BZSet A) Z (BZAction A a, BZEquiv A a)).1.1 @ -i\r
\r
-decodeZ1 (A : BZ) (a : BZSet A) (x : Z) : Id Z (transport (decodeZP A a) (BZS A (transport (<i> (decodeZP A a) @ -i) x)))\r
+decodeZ1 (A : BZ) (a : BZSet A) (x : Z) : Path Z (transport (decodeZP A a) (BZS A (transport (<i> (decodeZP A a) @ -i) x)))\r
(BZEquiv A a (BZS A (BZAction A a x))).1.1\r
- = compId Z (transport (decodeZP A a) (BZS A (transport (<i> (decodeZP A a) @ -i) x)))\r
+ = compPath Z (transport (decodeZP A a) (BZS A (transport (<i> (decodeZP A a) @ -i) x)))\r
(transport (decodeZP A a) (BZS A (BZAction A a x)))\r
(BZEquiv A a (BZS A (BZAction A a x))).1.1\r
(<i> transport (decodeZP A a) (BZS A (lem0 Z (BZSet A) (BZAction A a) (BZEquiv A a) x @ i)))\r
(lem1 Z (BZSet A) (BZAction A a) (BZEquiv A a) (BZS A (BZAction A a x)))\r
opaque decodeZ1\r
\r
-decodeZ2 (A : BZ) (a : BZSet A) (x : Z) : Id Z (BZEquiv A a (BZS A (BZAction A a x))).1.1 (sucZ x)\r
- = compId Z (BZEquiv A a (BZS A (BZAction A a x))).1.1 (BZEquiv A a (BZAction A a (sucZ x))).1.1 (sucZ x)\r
+decodeZ2 (A : BZ) (a : BZSet A) (x : Z) : Path Z (BZEquiv A a (BZS A (BZAction A a x))).1.1 (sucZ x)\r
+ = compPath Z (BZEquiv A a (BZS A (BZAction A a x))).1.1 (BZEquiv A a (BZAction A a (sucZ x))).1.1 (sucZ x)\r
(<i> (BZEquiv A a (lem2 (BZSet A) (BZShift A) a x @ i)).1.1)\r
(secEq Z (BZSet A) (BZAction A a, BZEquiv A a) (sucZ x))\r
opaque decodeZ2\r
\r
-decodeZ3 (A : BZ) (a : BZSet A) : Id (Z->Z) (\(x : Z) -> transport (decodeZP A a) (BZS A (transport (<i> (decodeZP A a) @ -i) x))) sucZ\r
- = <i> \(x : Z) -> compId Z (transport (decodeZP A a) (BZS A (transport (<i> (decodeZP A a) @ -i) x))) (BZEquiv A a (BZS A (BZAction A a x))).1.1 (sucZ x) (decodeZ1 A a x) (decodeZ2 A a x) @ i\r
+decodeZ3 (A : BZ) (a : BZSet A) : Path (Z->Z) (\(x : Z) -> transport (decodeZP A a) (BZS A (transport (<i> (decodeZP A a) @ -i) x))) sucZ\r
+ = <i> \(x : Z) -> compPath Z (transport (decodeZP A a) (BZS A (transport (<i> (decodeZP A a) @ -i) x))) (BZEquiv A a (BZS A (BZAction A a x))).1.1 (sucZ x) (decodeZ1 A a x) (decodeZ2 A a x) @ i\r
opaque decodeZ3\r
\r
-decodeZQ (A : BZ) (a : BZSet A) : IdP (<i> ((decodeZP A a)@-i) -> ((decodeZP A a)@-i)) sucZ (BZS A)\r
- = <i> substIdP (BZSet A -> BZSet A) (Z -> Z) (<i> ((decodeZP A a)@i) -> ((decodeZP A a)@i)) (BZS A) sucZ (decodeZ3 A a) @ -i\r
+decodeZQ (A : BZ) (a : BZSet A) : PathP (<i> ((decodeZP A a)@-i) -> ((decodeZP A a)@-i)) sucZ (BZS A)\r
+ = <i> substPathP (BZSet A -> BZSet A) (Z -> Z) (<i> ((decodeZP A a)@i) -> ((decodeZP A a)@i)) (BZS A) sucZ (decodeZ3 A a) @ -i\r
opaque decodeZQ\r
\r
-decodeZ (A : BZ) (a : BZSet A) : Id BZ ZBZ A = (lemBZ ZBZ A).1 (<i> decodeZP A a @ -i, decodeZQ A a)\r
+decodeZ (A : BZ) (a : BZSet A) : Path BZ ZBZ A = (lemBZ ZBZ A).1 (<i> decodeZP A a @ -i, decodeZQ A a)\r
\r
AA : U = (X : U) * (X -> X)\r
-BZ' : U = (X : AA) * ishinh_UU (Id AA X (Z, sucZ))\r
-BZequivBZ' : Id U BZ BZ' = isoId BZ BZ' F G FG GF\r
+BZ' : U = (X : AA) * ishinh_UU (Path AA X (Z, sucZ))\r
+BZequivBZ' : Path U BZ BZ' = isoPath BZ BZ' F G FG GF\r
where\r
- F (A : BZ) : BZ' = ((BZSet A, BZS A), BZNE A (ishinh (Id AA (BZSet A, BZS A) (Z, sucZ)))\r
- (\(a : BZSet A) -> hinhpr (Id AA (BZSet A, BZS A) (Z, sucZ)) (<i> (decodeZP A a @ i, decodeZQ A a @ -i))))\r
+ F (A : BZ) : BZ' = ((BZSet A, BZS A), BZNE A (ishinh (Path AA (BZSet A, BZS A) (Z, sucZ)))\r
+ (\(a : BZSet A) -> hinhpr (Path AA (BZSet A, BZS A) (Z, sucZ)) (<i> (decodeZP A a @ i, decodeZQ A a @ -i))))\r
G (A : BZ') : BZ = (A.1.1,\r
- (A.2 (ishinh A.1.1) (\(a : Id AA A.1 (Z, sucZ)) -> hinhpr A.1.1 (transport (<i> (a @ -i).1) zeroZ)),\r
+ (A.2 (ishinh A.1.1) (\(a : Path AA A.1 (Z, sucZ)) -> hinhpr A.1.1 (transport (<i> (a @ -i).1) zeroZ)),\r
((A.1.2, SHIFT), EQUIV)))\r
where\r
SHIFT : isEquiv A.1.1 A.1.1 A.1.2\r
= A.2 (isEquiv A.1.1 A.1.1 A.1.2, propIsEquiv A.1.1 A.1.1 A.1.2)\r
- (\(a : Id AA A.1 (Z, sucZ)) -> transport (<i> isEquiv (a@-i).1 (a@-i).1 (a@-i).2) (gradLemma Z Z sucZ predZ sucpredZ predsucZ))\r
+ (\(a : Path AA A.1 (Z, sucZ)) -> transport (<i> isEquiv (a@-i).1 (a@-i).1 (a@-i).2) (gradLemma Z Z sucZ predZ sucpredZ predsucZ))\r
ZEquiv : (x : Z) -> isEquiv Z Z (action Z (sucZ, gradLemma Z Z sucZ predZ sucpredZ predsucZ) x) = BZEquiv ZBZ\r
- hole (a : Id AA A.1 (Z, sucZ))\r
- : IdP (<i> isEquiv (a@-i).1 (a@-i).1 (a@-i).2) (gradLemma Z Z sucZ predZ sucpredZ predsucZ) SHIFT\r
- = lemIdPProp (isEquiv Z Z sucZ) (isEquiv A.1.1 A.1.1 A.1.2) (propIsEquiv Z Z sucZ)\r
+ hole (a : Path AA A.1 (Z, sucZ))\r
+ : PathP (<i> isEquiv (a@-i).1 (a@-i).1 (a@-i).2) (gradLemma Z Z sucZ predZ sucpredZ predsucZ) SHIFT\r
+ = lemPathPProp (isEquiv Z Z sucZ) (isEquiv A.1.1 A.1.1 A.1.2) (propIsEquiv Z Z sucZ)\r
(<i> isEquiv (a@-i).1 (a@-i).1 (a@-i).2) (gradLemma Z Z sucZ predZ sucpredZ predsucZ) SHIFT\r
- ZEquivEq (a : Id AA A.1 (Z, sucZ))\r
- : Id U ((x : Z) -> isEquiv Z Z (action Z (sucZ, gradLemma Z Z sucZ predZ sucpredZ predsucZ) x))\r
+ ZEquivEq (a : Path AA A.1 (Z, sucZ))\r
+ : Path U ((x : Z) -> isEquiv Z Z (action Z (sucZ, gradLemma Z Z sucZ predZ sucpredZ predsucZ) x))\r
((x : A.1.1) -> isEquiv Z A.1.1 (action A.1.1 (A.1.2, SHIFT) x))\r
= <i> (x : (a@-i).1) -> isEquiv Z (a@-i).1 (action (a@-i).1 ((a@-i).2, hole a @ i) x)\r
EQUIV : (x : A.1.1) -> isEquiv Z A.1.1 (action A.1.1 (A.1.2, SHIFT) x)\r
= A.2 ((x : A.1.1) -> isEquiv Z A.1.1 (action A.1.1 (A.1.2, SHIFT) x)\r
, propPi A.1.1 (\(x : A.1.1) -> isEquiv Z A.1.1 (action A.1.1 (A.1.2, SHIFT) x))\r
(\(x : A.1.1) -> propIsEquiv Z A.1.1 (action A.1.1 (A.1.2, SHIFT) x)))\r
- (\(a : Id AA A.1 (Z, sucZ)) -> transport (ZEquivEq a) ZEquiv)\r
- FG (A : BZ') : Id BZ' (F (G A)) A = <i> ((A.1.1, A.1.2), propishinh (Id AA (A.1.1, A.1.2) (Z, sucZ)) (F (G A)).2 A.2 @ i)\r
- GF (A : BZ) : Id BZ (G (F A)) A = (lemBZ (G (F A)) A).1 (<_> BZSet A, <_> BZS A)\r
+ (\(a : Path AA A.1 (Z, sucZ)) -> transport (ZEquivEq a) ZEquiv)\r
+ FG (A : BZ') : Path BZ' (F (G A)) A = <i> ((A.1.1, A.1.2), propishinh (Path AA (A.1.1, A.1.2) (Z, sucZ)) (F (G A)).2 A.2 @ i)\r
+ GF (A : BZ) : Path BZ (G (F A)) A = (lemBZ (G (F A)) A).1 (<_> BZSet A, <_> BZS A)\r
\r
-prf : Id (equiv Z Z) (BZAction ZBZ zeroZ, BZEquiv ZBZ zeroZ) (\(x : Z) -> x, isEquivId Z) = hole\r
+prf : Path (equiv Z Z) (BZAction ZBZ zeroZ, BZEquiv ZBZ zeroZ) (\(x : Z) -> x, isEquivPath Z) = hole\r
where\r
- hole0 : (x : Z) -> Id Z (BZAction ZBZ zeroZ x) x = split\r
+ hole0 : (x : Z) -> Path Z (BZAction ZBZ zeroZ x) x = split\r
inl n -> hole1 n\r
where\r
- hole1 : (n : nat) -> Id Z (BZAction ZBZ zeroZ (inl n)) (inl n) = split\r
+ hole1 : (n : nat) -> Path Z (BZAction ZBZ zeroZ (inl n)) (inl n) = split\r
zero -> <_> inl zero\r
suc n -> mapOnPath Z Z predZ (BZAction ZBZ zeroZ (inl n)) (inl n) (hole1 n)\r
inr n -> hole1 n\r
where\r
- hole1 : (n : nat) -> Id Z (BZAction ZBZ zeroZ (inr n)) (inr n) = split\r
+ hole1 : (n : nat) -> Path Z (BZAction ZBZ zeroZ (inr n)) (inr n) = split\r
zero -> <_> inr zero\r
suc n -> mapOnPath Z Z sucZ (BZAction ZBZ zeroZ (inr n)) (inr n) (hole1 n)\r
- hole : Id (equiv Z Z) (BZAction ZBZ zeroZ, BZEquiv ZBZ zeroZ) (\(x : Z) -> x, isEquivId Z)\r
+ hole : Path (equiv Z Z) (BZAction ZBZ zeroZ, BZEquiv ZBZ zeroZ) (\(x : Z) -> x, isEquivPath Z)\r
= <i> (\(x : Z) -> hole0 x @ i,\r
- lemIdPProp\r
+ lemPathPProp\r
(isEquiv Z Z (BZAction ZBZ zeroZ)) (isEquiv Z Z (\(x : Z) -> x))\r
(propIsEquiv Z Z (BZAction ZBZ zeroZ))\r
(<j> isEquiv Z Z (\(x : Z) -> hole0 x @ j))\r
- (BZEquiv ZBZ zeroZ) (isEquivId Z) @ i\r
+ (BZEquiv ZBZ zeroZ) (isEquivPath Z) @ i\r
)\r
opaque prf\r
\r
decodeEncodeZRefl0\r
- : Id (Id U Z Z) (univalence Z Z (BZAction ZBZ zeroZ, BZEquiv ZBZ zeroZ)).1.1 (<_> Z)\r
- = transport (<i> Id (Id U Z Z) (univalence Z Z (prf @ -i)).1.1 (<_> Z)) (lem11 Z)\r
+ : Path (Path U Z Z) (univalence Z Z (BZAction ZBZ zeroZ, BZEquiv ZBZ zeroZ)).1.1 (<_> Z)\r
+ = transport (<i> Path (Path U Z Z) (univalence Z Z (prf @ -i)).1.1 (<_> Z)) (lem11 Z)\r
opaque decodeEncodeZRefl0\r
\r
decodeEncodeZRefl1\r
- : IdP (<j> (IdP (<i> decodeEncodeZRefl0@j@i -> decodeEncodeZRefl0@j@i) sucZ sucZ))\r
+ : PathP (<j> (PathP (<i> decodeEncodeZRefl0@j@i -> decodeEncodeZRefl0@j@i) sucZ sucZ))\r
(<i> (BZShift (decodeZ ZBZ zeroZ@i)).1) (<_> sucZ)\r
- = lemIdPSet2 (Z->Z) (Z->Z) (setPi Z (\(_ : Z) -> Z) (\(_ : Z) -> ZSet))\r
+ = lemPathPSet2 (Z->Z) (Z->Z) (setPi Z (\(_ : Z) -> Z) (\(_ : Z) -> ZSet))\r
(<j> decodeEncodeZRefl0@0@j -> decodeEncodeZRefl0@0@j)\r
(<j> decodeEncodeZRefl0@1@j -> decodeEncodeZRefl0@1@j)\r
(<i j> decodeEncodeZRefl0@i@j -> decodeEncodeZRefl0@i@j)\r
sucZ sucZ (<i> (BZShift (decodeZ ZBZ zeroZ@i)).1) (<_> sucZ)\r
opaque decodeEncodeZRefl1\r
\r
-decodeEncodeZRefl2 : Id ((p : Id U Z Z) * IdP (<i> p@i -> p@i) sucZ sucZ) ((lemBZ ZBZ ZBZ).2 (decodeZ ZBZ zeroZ)).1.1 ((lemBZ ZBZ ZBZ).2 (<_> ZBZ)).1.1\r
+decodeEncodeZRefl2 : Path ((p : Path U Z Z) * PathP (<i> p@i -> p@i) sucZ sucZ) ((lemBZ ZBZ ZBZ).2 (decodeZ ZBZ zeroZ)).1.1 ((lemBZ ZBZ ZBZ).2 (<_> ZBZ)).1.1\r
= <i> (decodeEncodeZRefl0 @ i, decodeEncodeZRefl1 @ i)\r
opaque decodeEncodeZRefl2\r
\r
\r
-decodeEncodeZRefl : Id loopBZ (decodeZ ZBZ zeroZ) (<_> ZBZ)\r
- = lem10 ((p : Id U Z Z) * IdP (<i> p@i -> p@i) sucZ sucZ) loopBZ (lemBZ ZBZ ZBZ) (decodeZ ZBZ zeroZ) (<_> ZBZ) decodeEncodeZRefl2\r
+decodeEncodeZRefl : Path loopBZ (decodeZ ZBZ zeroZ) (<_> ZBZ)\r
+ = lem10 ((p : Path U Z Z) * PathP (<i> p@i -> p@i) sucZ sucZ) loopBZ (lemBZ ZBZ ZBZ) (decodeZ ZBZ zeroZ) (<_> ZBZ) decodeEncodeZRefl2\r
opaque decodeEncodeZRefl\r
\r
-decodeEncodeZ : (A : BZ) -> (p : Id BZ ZBZ A) -> Id (Id BZ ZBZ A) (decodeZ A (encodeZ A p)) p\r
- = J BZ ZBZ (\(A : BZ) -> \(p : Id BZ ZBZ A) -> Id (Id BZ ZBZ A) (decodeZ A (encodeZ A p)) p) decodeEncodeZRefl\r
+decodeEncodeZ : (A : BZ) -> (p : Path BZ ZBZ A) -> Path (Path BZ ZBZ A) (decodeZ A (encodeZ A p)) p\r
+ = J BZ ZBZ (\(A : BZ) -> \(p : Path BZ ZBZ A) -> Path (Path BZ ZBZ A) (decodeZ A (encodeZ A p)) p) decodeEncodeZRefl\r
opaque decodeEncodeZ\r
\r
-encodeDecodeZ (A : BZ) (p : BZSet A) : Id (BZSet A) (transport (univalence (BZSet A) Z (BZAction A p, BZEquiv A p)).1.1 zeroZ) p\r
+encodeDecodeZ (A : BZ) (p : BZSet A) : Path (BZSet A) (transport (univalence (BZSet A) Z (BZAction A p, BZEquiv A p)).1.1 zeroZ) p\r
= lem0 Z (BZSet A) (BZAction A p) (BZEquiv A p) zeroZ\r
opaque encodeDecodeZ\r
\r
-encodeLoop (x : loopBZ) : Id Z (encodeZ ZBZ (compBZ x loopZ)) (sucZ (encodeZ ZBZ x))\r
- = J BZ ZBZ (\(A : BZ) -> \(x : Id BZ ZBZ A) -> Id (BZSet A) (encodeZ A (compId BZ ZBZ A A x (loopA A))) (BZS A (encodeZ A x)))\r
+encodeLoop (x : loopBZ) : Path Z (encodeZ ZBZ (compBZ x loopZ)) (sucZ (encodeZ ZBZ x))\r
+ = J BZ ZBZ (\(A : BZ) -> \(x : Path BZ ZBZ A) -> Path (BZSet A) (encodeZ A (compPath BZ ZBZ A A x (loopA A))) (BZS A (encodeZ A x)))\r
(<_> inr (suc zero)) ZBZ x\r
opaque encodeLoop\r
\r
-encodeInvLoop (x : loopBZ) : Id Z (encodeZ ZBZ (compBZ x invLoopZ)) (predZ (encodeZ ZBZ x))\r
- = J BZ ZBZ (\(A : BZ) -> \(x : Id BZ ZBZ A) -> Id (BZSet A) (encodeZ A (compId BZ ZBZ A A x (<i> loopA A @ -i))) (BZP A (encodeZ A x)))\r
+encodeInvLoop (x : loopBZ) : Path Z (encodeZ ZBZ (compBZ x invLoopZ)) (predZ (encodeZ ZBZ x))\r
+ = J BZ ZBZ (\(A : BZ) -> \(x : Path BZ ZBZ A) -> Path (BZSet A) (encodeZ A (compPath BZ ZBZ A A x (<i> loopA A @ -i))) (BZP A (encodeZ A x)))\r
(<_> inl zero) ZBZ x\r
opaque encodeInvLoop\r
\r
-decodeLoop (z : Z) : Id loopBZ (compBZ (decodeZ ZBZ z) loopZ) (decodeZ ZBZ (sucZ z))\r
- = transport (<i> Id loopBZ (decodeEncodeZ ZBZ (compBZ (decodeZ ZBZ z) loopZ) @ i) (decodeZ ZBZ (sucZ (encodeDecodeZ ZBZ z @ i))))\r
+decodeLoop (z : Z) : Path loopBZ (compBZ (decodeZ ZBZ z) loopZ) (decodeZ ZBZ (sucZ z))\r
+ = transport (<i> Path loopBZ (decodeEncodeZ ZBZ (compBZ (decodeZ ZBZ z) loopZ) @ i) (decodeZ ZBZ (sucZ (encodeDecodeZ ZBZ z @ i))))\r
(<i> decodeZ ZBZ (encodeLoop (decodeZ ZBZ z) @ i))\r
opaque decodeLoop\r
-decodeInvLoop (z : Z) : Id (Id BZ ZBZ ZBZ) (compBZ (decodeZ ZBZ z) invLoopZ) (decodeZ ZBZ (predZ z))\r
- = transport (<i> Id loopBZ (decodeEncodeZ ZBZ (compBZ (decodeZ ZBZ z) invLoopZ) @ i) (decodeZ ZBZ (predZ (encodeDecodeZ ZBZ z @ i))))\r
+decodeInvLoop (z : Z) : Path (Path BZ ZBZ ZBZ) (compBZ (decodeZ ZBZ z) invLoopZ) (decodeZ ZBZ (predZ z))\r
+ = transport (<i> Path loopBZ (decodeEncodeZ ZBZ (compBZ (decodeZ ZBZ z) invLoopZ) @ i) (decodeZ ZBZ (predZ (encodeDecodeZ ZBZ z @ i))))\r
(<i> decodeZ ZBZ (encodeInvLoop (decodeZ ZBZ z) @ i))\r
opaque decodeInvLoop\r
\r
where\r
opaque decodeZ\r
opaque loopA\r
- T : U = Id BZ A B\r
+ T : U = Path BZ A B\r
TNE : ishinh_UU T = BZNE A (ishinh T) (\(a : BZSet A) -> BZNE B (ishinh T) (\(b : BZSet B) ->\r
- hinhpr T (compId BZ A ZBZ B (<i> decodeZ A a @ -i) (decodeZ B b))))\r
- F (A B : BZ) (x : (Id BZ A B)) : (Id BZ A B) = compId BZ A B B x (loopA B)\r
- G (A B : BZ) (x : (Id BZ A B)) : (Id BZ A B) = compId BZ A B B x (<i> loopA B @ -i)\r
- FG (A B : BZ) (x : (Id BZ A B)) : Id (Id BZ A B) (F A B (G A B x)) x = lemCompInv BZ A B B x (<i> loopA B @ -i)\r
+ hinhpr T (compPath BZ A ZBZ B (<i> decodeZ A a @ -i) (decodeZ B b))))\r
+ F (A B : BZ) (x : (Path BZ A B)) : (Path BZ A B) = compPath BZ A B B x (loopA B)\r
+ G (A B : BZ) (x : (Path BZ A B)) : (Path BZ A B) = compPath BZ A B B x (<i> loopA B @ -i)\r
+ FG (A B : BZ) (x : (Path BZ A B)) : Path (Path BZ A B) (F A B (G A B x)) x = lemCompInv BZ A B B x (<i> loopA B @ -i)\r
opaque FG\r
- GF (A B : BZ) (x : (Id BZ A B)) : Id (Id BZ A B) (G A B (F A B x)) x = lemCompInv BZ A B B x (loopA B)\r
+ GF (A B : BZ) (x : (Path BZ A B)) : Path (Path BZ A B) (G A B (F A B x)) x = lemCompInv BZ A B B x (loopA B)\r
opaque GF\r
- TShift (A B : BZ) : equiv (Id BZ A B) (Id BZ A B) = (F A B, gradLemma (Id BZ A B) (Id BZ A B) (F A B) (G A B) (FG A B) (GF A B))\r
- hole : (y : Z) -> Id (Id BZ ZBZ ZBZ) (action (Id BZ ZBZ ZBZ) (TShift ZBZ ZBZ) (<_> ZBZ) y) (decodeZ ZBZ y) = split\r
+ TShift (A B : BZ) : equiv (Path BZ A B) (Path BZ A B) = (F A B, gradLemma (Path BZ A B) (Path BZ A B) (F A B) (G A B) (FG A B) (GF A B))\r
+ hole : (y : Z) -> Path (Path BZ ZBZ ZBZ) (action (Path BZ ZBZ ZBZ) (TShift ZBZ ZBZ) (<_> ZBZ) y) (decodeZ ZBZ y) = split\r
inl u -> hole0 u\r
- where hole0 : (n : nat) -> Id (Id BZ ZBZ ZBZ) (action (Id BZ ZBZ ZBZ) (TShift ZBZ ZBZ) (<_> ZBZ) (inl n)) (decodeZ ZBZ (inl n)) = split\r
+ where hole0 : (n : nat) -> Path (Path BZ ZBZ ZBZ) (action (Path BZ ZBZ ZBZ) (TShift ZBZ ZBZ) (<_> ZBZ) (inl n)) (decodeZ ZBZ (inl n)) = split\r
zero -> hole1\r
- where hole1 : Id (Id BZ ZBZ ZBZ) (compBZ (<_> ZBZ) invLoopZ) (decodeZ ZBZ (inl zero))\r
- = compId loopBZ (compBZ (<_> ZBZ) invLoopZ) (compBZ (decodeZ ZBZ (inr zero)) invLoopZ) (decodeZ ZBZ (inl zero))\r
+ where hole1 : Path (Path BZ ZBZ ZBZ) (compBZ (<_> ZBZ) invLoopZ) (decodeZ ZBZ (inl zero))\r
+ = compPath loopBZ (compBZ (<_> ZBZ) invLoopZ) (compBZ (decodeZ ZBZ (inr zero)) invLoopZ) (decodeZ ZBZ (inl zero))\r
(<i> compBZ (decodeEncodeZRefl @ -i) invLoopZ)\r
(decodeInvLoop (inr zero))\r
suc n -> hole1\r
- where hole1 : Id (Id BZ ZBZ ZBZ) (compBZ (action loopBZ (TShift ZBZ ZBZ) (<_> ZBZ) (inl n)) invLoopZ) (decodeZ ZBZ (inl (suc n)))\r
- = compId (Id BZ ZBZ ZBZ) (compBZ (action (Id BZ ZBZ ZBZ) (TShift ZBZ ZBZ) (<_> ZBZ) (inl n)) invLoopZ)\r
+ where hole1 : Path (Path BZ ZBZ ZBZ) (compBZ (action loopBZ (TShift ZBZ ZBZ) (<_> ZBZ) (inl n)) invLoopZ) (decodeZ ZBZ (inl (suc n)))\r
+ = compPath (Path BZ ZBZ ZBZ) (compBZ (action (Path BZ ZBZ ZBZ) (TShift ZBZ ZBZ) (<_> ZBZ) (inl n)) invLoopZ)\r
(compBZ (decodeZ ZBZ (inl n)) invLoopZ) (decodeZ ZBZ (inl (suc n)))\r
(<i> compBZ (hole0 n @ i) invLoopZ)\r
(decodeInvLoop (inl n))\r
inr u -> hole0 u\r
- where hole0 : (n : nat) -> Id (Id BZ ZBZ ZBZ) (action (Id BZ ZBZ ZBZ) (TShift ZBZ ZBZ) (<_> ZBZ) (inr n)) (decodeZ ZBZ (inr n)) = split\r
+ where hole0 : (n : nat) -> Path (Path BZ ZBZ ZBZ) (action (Path BZ ZBZ ZBZ) (TShift ZBZ ZBZ) (<_> ZBZ) (inr n)) (decodeZ ZBZ (inr n)) = split\r
zero -> <i> decodeEncodeZRefl @ -i\r
- suc n -> compId (Id BZ ZBZ ZBZ) (compBZ (action (Id BZ ZBZ ZBZ) (TShift ZBZ ZBZ) (<_> ZBZ) (inr n)) loopZ)\r
+ suc n -> compPath (Path BZ ZBZ ZBZ) (compBZ (action (Path BZ ZBZ ZBZ) (TShift ZBZ ZBZ) (<_> ZBZ) (inr n)) loopZ)\r
(compBZ (decodeZ ZBZ (inr n)) loopZ) (decodeZ ZBZ (inr (suc n)))\r
(<i> compBZ (hole0 n @ i) loopZ)\r
(decodeLoop (inr n))\r
visible decodeZ\r
- TEquiv''' : isEquiv Z (Id BZ ZBZ ZBZ) (action (Id BZ ZBZ ZBZ) (TShift ZBZ ZBZ) (<_> ZBZ))\r
+ TEquiv''' : isEquiv Z (Path BZ ZBZ ZBZ) (action (Path BZ ZBZ ZBZ) (TShift ZBZ ZBZ) (<_> ZBZ))\r
= transport (<i> isEquiv Z loopBZ (\(y : Z) -> hole y @ -i)) (gradLemma Z loopBZ (decodeZ ZBZ) (encodeZ ZBZ) (decodeEncodeZ ZBZ) (encodeDecodeZ ZBZ))\r
- TEquiv'' (b : BZSet B) (x : Id BZ ZBZ B) : isEquiv Z (Id BZ ZBZ B) (action (Id BZ ZBZ B) (TShift ZBZ B) x)\r
- = J BZ ZBZ (\(B : BZ) -> \(x : Id BZ ZBZ B) -> isEquiv Z (Id BZ ZBZ B) (action (Id BZ ZBZ B) (TShift ZBZ B) x))\r
+ TEquiv'' (b : BZSet B) (x : Path BZ ZBZ B) : isEquiv Z (Path BZ ZBZ B) (action (Path BZ ZBZ B) (TShift ZBZ B) x)\r
+ = J BZ ZBZ (\(B : BZ) -> \(x : Path BZ ZBZ B) -> isEquiv Z (Path BZ ZBZ B) (action (Path BZ ZBZ B) (TShift ZBZ B) x))\r
TEquiv''' B x\r
TEquiv' (a : BZSet A) (b : BZSet B) : (x : T) -> isEquiv Z T (action T (TShift A B) x)\r
- = J BZ ZBZ (\(A : BZ) -> \(p : Id BZ ZBZ A) -> (x : Id BZ A B) -> isEquiv Z (Id BZ A B) (action (Id BZ A B) (TShift A B) x))\r
+ = J BZ ZBZ (\(A : BZ) -> \(p : Path BZ ZBZ A) -> (x : Path BZ A B) -> isEquiv Z (Path BZ A B) (action (Path BZ A B) (TShift A B) x))\r
(TEquiv'' b) A (decodeZ A a)\r
TEquiv (x : T) : isEquiv Z T (action T (TShift A B) x)\r
= BZNE A (isEquiv Z T (action T (TShift A B) x), propIsEquiv Z T (action T (TShift A B) x))\r
(\(b : BZSet B) -> TEquiv' a b x))\r
visible decodeZ\r
\r
-loopBZequalsZ : Id U loopBZ Z = isoId loopBZ Z (encodeZ ZBZ) (decodeZ ZBZ) (encodeDecodeZ ZBZ) (decodeEncodeZ ZBZ)\r
+loopBZequalsZ : Path U loopBZ Z = isoPath loopBZ Z (encodeZ ZBZ) (decodeZ ZBZ) (encodeDecodeZ ZBZ) (decodeEncodeZ ZBZ)\r
\r
-loopS1equalsLoopBZ : Id U loopS1 loopBZ = compId U loopS1 Z loopBZ loopS1equalsZ (<i> loopBZequalsZ @ -i)\r
+loopS1equalsLoopBZ : Path U loopS1 loopBZ = compPath U loopS1 Z loopBZ loopS1equalsZ (<i> loopBZequalsZ @ -i)\r
loopS1equalsLoopBZ' : equiv loopS1 loopBZ = transEquiv' loopBZ loopS1 loopS1equalsLoopBZ\r
\r
-- BZ = S1\r
\r
-- mapOnPath S1 BZ F base base = loopS1equalsLoopBZ'.1 : same left inverse and loopS1equalsLoopBZ'.1 is an equivalence\r
\r
-G : (y : S1) -> Id BZ ZBZ (F y) -> Id S1 base y = split\r
+G : (y : S1) -> Path BZ ZBZ (F y) -> Path S1 base y = split\r
base -> \(x : loopBZ) -> (loopS1equalsLoopBZ'.2 x).1.1\r
loop @ i -> hole @ i\r
where\r
- hole4 (x : Z) : Id loopS1 (loopIt (predZ x)) (compS1 (loopIt x) invLoop)\r
+ hole4 (x : Z) : Path loopS1 (loopIt (predZ x)) (compS1 (loopIt x) invLoop)\r
= lem2It x\r
hole3 (x : loopBZ)\r
- : Id loopS1\r
+ : Path loopS1\r
(compS1 (decode base (encodeZ ZBZ (compBZ x invLoopZ))) loop1)\r
(decode base (encodeZ ZBZ x))\r
- = compId loopS1\r
+ = compPath loopS1\r
(compS1 (decode base (encodeZ ZBZ (compBZ x invLoopZ))) loop1)\r
(compS1 (decode base (predZ (encodeZ ZBZ x))) loop1)\r
(decode base (encodeZ ZBZ x))\r
(<i> compS1 (decode base (encodeInvLoop x @ i)) loop1)\r
- (compId loopS1\r
+ (compPath loopS1\r
(compS1 (decode base (predZ (encodeZ ZBZ x))) loop1)\r
(compS1 (compS1 (decode base (encodeZ ZBZ x)) invLoop) loop1)\r
(decode base (encodeZ ZBZ x))\r
(<i> compS1 (hole4 (encodeZ ZBZ x) @ i) loop1)\r
(<i> compInv S1 base base (decode base (encodeZ ZBZ x)) base invLoop @ -i))\r
- hole6 (x : loopBZ) : Id loopS1 (loopS1equalsLoopBZ'.2 x).1.1 (decode base (encodeZ ZBZ x))\r
- = compId loopS1 (loopS1equalsLoopBZ'.2 x).1.1 (loopIt (transConstN Z (encodeZ ZBZ x) (suc (suc (suc (suc (suc zero))))))) (decode base (encodeZ ZBZ x))\r
+ hole6 (x : loopBZ) : Path loopS1 (loopS1equalsLoopBZ'.2 x).1.1 (decode base (encodeZ ZBZ x))\r
+ = compPath loopS1 (loopS1equalsLoopBZ'.2 x).1.1 (loopIt (transConstN Z (encodeZ ZBZ x) (suc (suc (suc (suc (suc zero))))))) (decode base (encodeZ ZBZ x))\r
(lemHcomp3 (loopIt (transConstN Z (encodeZ ZBZ x) (suc (suc (suc (suc (suc zero))))))))\r
(<i> decode base (transConstNRefl Z (encodeZ ZBZ x) (suc (suc (suc (suc (suc zero))))) @ i))\r
- hole1 : Id (loopBZ -> loopS1)\r
+ hole1 : Path (loopBZ -> loopS1)\r
(\(x : loopBZ) -> compS1 ((loopS1equalsLoopBZ'.2 (compBZ x invLoopZ)).1.1) loop1)\r
(\(x : loopBZ) -> (loopS1equalsLoopBZ'.2 x).1.1)\r
= <j> \(x : loopBZ) ->\r
- transport (<i> Id loopS1 (compS1 (hole6 (compBZ x invLoopZ) @ -i) loop1) (loopS1equalsLoopBZ'.2 x).1.1)\r
- (transport (<i> Id loopS1 (compS1 (decode base (encodeZ ZBZ (compBZ x invLoopZ))) loop1) (hole6 x @ -i))\r
+ transport (<i> Path loopS1 (compS1 (hole6 (compBZ x invLoopZ) @ -i) loop1) (loopS1equalsLoopBZ'.2 x).1.1)\r
+ (transport (<i> Path loopS1 (compS1 (decode base (encodeZ ZBZ (compBZ x invLoopZ))) loop1) (hole6 x @ -i))\r
(hole3 x)) @ j\r
- hole : IdP (<i> Id BZ ZBZ (F (loop1 @ i)) -> Id S1 base (loop1 @ i))\r
+ hole : PathP (<i> Path BZ ZBZ (F (loop1 @ i)) -> Path S1 base (loop1 @ i))\r
(\(x : loopBZ) -> (loopS1equalsLoopBZ'.2 x).1.1)\r
(\(x : loopBZ) -> (loopS1equalsLoopBZ'.2 x).1.1)\r
- = substIdP (loopBZ -> loopS1) (loopBZ -> loopS1)\r
- (<i> Id BZ ZBZ (F (loop1 @ i)) -> Id S1 base (loop1 @ i))\r
+ = substPathP (loopBZ -> loopS1) (loopBZ -> loopS1)\r
+ (<i> Path BZ ZBZ (F (loop1 @ i)) -> Path S1 base (loop1 @ i))\r
(\(x : loopBZ) -> (loopS1equalsLoopBZ'.2 x).1.1)\r
(\(x : loopBZ) -> (loopS1equalsLoopBZ'.2 x).1.1)\r
hole1\r
\r
-GF : (y : S1) -> (x : Id S1 base y) -> Id (Id S1 base y) (G y (mapOnPath S1 BZ F base y x)) x\r
- = J S1 base (\(y : S1) -> \(x : Id S1 base y) -> Id (Id S1 base y) (G y (mapOnPath S1 BZ F base y x)) x)\r
+GF : (y : S1) -> (x : Path S1 base y) -> Path (Path S1 base y) (G y (mapOnPath S1 BZ F base y x)) x\r
+ = J S1 base (\(y : S1) -> \(x : Path S1 base y) -> Path (Path S1 base y) (G y (mapOnPath S1 BZ F base y x)) x)\r
(lemHcomp3 (<_> base))\r
opaque GF\r
\r
-F_fullyFaithful0 : Id (loopS1 -> loopBZ) (mapOnPath S1 BZ F base base) loopS1equalsLoopBZ'.1\r
+F_fullyFaithful0 : Path (loopS1 -> loopBZ) (mapOnPath S1 BZ F base base) loopS1equalsLoopBZ'.1\r
= lemEquiv1 loopS1 loopBZ (mapOnPath S1 BZ F base base) loopS1equalsLoopBZ'.1 loopS1equalsLoopBZ'.2 (GF base)\r
opaque F_fullyFaithful0\r
\r
\r
-F_fullyFaithful : (x y : S1) -> isEquiv (Id S1 x y) (Id BZ (F x) (F y)) (mapOnPath S1 BZ F x y)\r
- = lemPropFib (\(x : S1) -> (y : S1) -> isEquiv (Id S1 x y) (Id BZ (F x) (F y)) (mapOnPath S1 BZ F x y))\r
- (\(x : S1) -> propPi S1 (\(y : S1) -> isEquiv (Id S1 x y) (Id BZ (F x) (F y)) (mapOnPath S1 BZ F x y))\r
- (\(y : S1) -> propIsEquiv (Id S1 x y) (Id BZ (F x) (F y)) (mapOnPath S1 BZ F x y)))\r
- (lemPropFib (\(y : S1) -> isEquiv (Id S1 base y) (Id BZ ZBZ (F y)) (mapOnPath S1 BZ F base y))\r
- (\(y : S1) -> propIsEquiv (Id S1 base y) (Id BZ ZBZ (F y)) (mapOnPath S1 BZ F base y))\r
+F_fullyFaithful : (x y : S1) -> isEquiv (Path S1 x y) (Path BZ (F x) (F y)) (mapOnPath S1 BZ F x y)\r
+ = lemPropFib (\(x : S1) -> (y : S1) -> isEquiv (Path S1 x y) (Path BZ (F x) (F y)) (mapOnPath S1 BZ F x y))\r
+ (\(x : S1) -> propPi S1 (\(y : S1) -> isEquiv (Path S1 x y) (Path BZ (F x) (F y)) (mapOnPath S1 BZ F x y))\r
+ (\(y : S1) -> propIsEquiv (Path S1 x y) (Path BZ (F x) (F y)) (mapOnPath S1 BZ F x y)))\r
+ (lemPropFib (\(y : S1) -> isEquiv (Path S1 base y) (Path BZ ZBZ (F y)) (mapOnPath S1 BZ F base y))\r
+ (\(y : S1) -> propIsEquiv (Path S1 base y) (Path BZ ZBZ (F y)) (mapOnPath S1 BZ F base y))\r
hole)\r
where\r
hole : isEquiv loopS1 loopBZ (mapOnPath S1 BZ F base base)\r
= transport (<i> isEquiv loopS1 loopBZ (F_fullyFaithful0 @ -i)) loopS1equalsLoopBZ'.2\r
opaque F_fullyFaithful\r
\r
-F_essentiallySurjective (y : BZ) : (x : S1) * Id BZ y (F x) = hole\r
+F_essentiallySurjective (y : BZ) : (x : S1) * Path BZ y (F x) = hole\r
where\r
- hInh (y : BZ) : ishinh_UU ((x : S1) * Id BZ y (F x)) = hole\r
+ hInh (y : BZ) : ishinh_UU ((x : S1) * Path BZ y (F x)) = hole\r
where\r
- hole2 (a : BZSet y) : (x : S1) * Id BZ y (F x) = (base, <i> decodeZ y a @ -i)\r
- hole1 (a : BZSet y) : ishinh_UU ((x : S1) * Id BZ y (F x)) = hinhpr ((x : S1) * Id BZ y (F x)) (hole2 a)\r
- hole : ishinh_UU ((x : S1) * Id BZ y (F x)) = BZNE y (ishinh_UU ((x : S1) * Id BZ y (F x)), propishinh ((x : S1) * Id BZ y (F x))) hole1\r
- hProp : prop ((x : S1) * Id BZ y (F x)) = transport (E13 S1 BZ F) (F_fullyFaithful) y\r
- hContr : isContr ((x : S1) * Id BZ y (F x)) = inhPropContr ((x : S1) * Id BZ y (F x)) hProp (hInh y)\r
- hole : (x : S1) * Id BZ y (F x) = hContr.1\r
-\r
-S1equalsBZ : Id U S1 BZ = hole\r
+ hole2 (a : BZSet y) : (x : S1) * Path BZ y (F x) = (base, <i> decodeZ y a @ -i)\r
+ hole1 (a : BZSet y) : ishinh_UU ((x : S1) * Path BZ y (F x)) = hinhpr ((x : S1) * Path BZ y (F x)) (hole2 a)\r
+ hole : ishinh_UU ((x : S1) * Path BZ y (F x)) = BZNE y (ishinh_UU ((x : S1) * Path BZ y (F x)), propishinh ((x : S1) * Path BZ y (F x))) hole1\r
+ hProp : prop ((x : S1) * Path BZ y (F x)) = transport (E13 S1 BZ F) (F_fullyFaithful) y\r
+ hContr : isContr ((x : S1) * Path BZ y (F x)) = inhPropContr ((x : S1) * Path BZ y (F x)) hProp (hInh y)\r
+ hole : (x : S1) * Path BZ y (F x) = hContr.1\r
+\r
+S1equalsBZ : Path U S1 BZ = hole\r
where\r
G (y : BZ) : S1 = (F_essentiallySurjective y).1\r
- FG (y : BZ) : Id BZ (F (G y)) y = <i> (F_essentiallySurjective y).2 @ -i\r
- GF (x : S1) : Id S1 (G (F x)) x = (F_fullyFaithful (G (F x)) x (FG (F x))).1.1\r
- hole : Id U S1 BZ = isoId S1 BZ F G FG GF\r
+ FG (y : BZ) : Path BZ (F (G y)) y = <i> (F_essentiallySurjective y).2 @ -i\r
+ GF (x : S1) : Path S1 (G (F x)) x = (F_fullyFaithful (G (F x)) x (FG (F x))).1.1\r
+ hole : Path U S1 BZ = isoPath S1 BZ F G FG GF\r
\r
invBZ : BZ -> BZ = transport (<i> S1equalsBZ@i -> S1equalsBZ@i) invS1\r
doubleLoopBZ : loopBZ -> loopBZ = transport (<i> loopS1equalsLoopBZ@i -> loopS1equalsLoopBZ@i) doubleLoop\r
-- EVAL: inr (suc (suc zero))\r
-- Time: 0m25.191s\r
\r
--- > let visible_all in let doubleLoop : Id BZ (multZ ZBZ ZBZ) (multZ ZBZ ZBZ) = <i> multZ (loopZ@-i) (loopZ@i) in transport (<j> BZSet (transport (<i> BZSet (doubleLoop @ i)) (<_> ZBZ) @ j)) zeroZ\r
+-- > let visible_all in let doubleLoop : Path BZ (multZ ZBZ ZBZ) (multZ ZBZ ZBZ) = <i> multZ (loopZ@-i) (loopZ@i) in transport (<j> BZSet (transport (<i> BZSet (doubleLoop @ i)) (<_> ZBZ) @ j)) zeroZ\r
-- too slow\r
\r
-- > let visible_all in transport (<j> BZSet (transport (<i> BZSet (multZ (loopZ@-i) (loopZ@i))) (<_> ZBZ) @ j)) zeroZ\r
loop @ x -> pathTwoT{Torus} @ x
-- branch for loop
-c2t_loop' : (c : S1) -> IdP (<_>Torus) (c2t_base c) (c2t_base c) = split
+c2t_loop' : (c : S1) -> PathP (<_>Torus) (c2t_base c) (c2t_base c) = split
base -> < x > pathOneT{Torus} @ x
loop @ y -> < x > squareT{Torus} @ y @ x
-- use funext to exchange the interval variable and the S1 variable
-c2t_loop : IdP (<_>S1 -> Torus) c2t_base c2t_base =
+c2t_loop : PathP (<_>S1 -> Torus) c2t_base c2t_base =
<y> \(c : S1) -> c2t_loop' c @ y
c2t' : S1 -> S1 -> Torus = split
------------------------------------------------------------------------
-- first composite: induct and reflexivity!
-t2c2t : (t : Torus) -> IdP (<_> Torus) (c2t (t2c t)) t = split
+t2c2t : (t : Torus) -> PathP (<_> Torus) (c2t (t2c t)) t = split
ptT -> <_> ptT
pathOneT @ y -> <_> pathOneT{Torus} @ y
pathTwoT @ x -> <_> pathTwoT{Torus} @ x
-- except we need to use the same tricks as in c2t to do the inner
-- induction
-c2t2c_base : (c2 : S1) -> IdP (<_> and S1 S1) (t2c (c2t_base c2)) (base,c2) = split
+c2t2c_base : (c2 : S1) -> PathP (<_> and S1 S1) (t2c (c2t_base c2)) (base,c2) = split
base -> <_> (base,base)
loop @ y -> <_> (base,loop1 @ y)
-- two binders exchanged.
-- c2t' (loop @ y) c2 = (c2t_loop @ y c2) = c2t_loop' c2 @ y
c2t2c_loop' : (c2 : S1) ->
- IdP (<y> IdP (<_> and S1 S1) (t2c (c2t_loop @ y c2)) (loop1 @ y , c2))
+ PathP (<y> PathP (<_> and S1 S1) (t2c (c2t_loop @ y c2)) (loop1 @ y , c2))
(c2t2c_base c2)
(c2t2c_base c2) = split
base -> <y> <_> (loop1 @ y, base)
loop @ x -> <y> <_> (loop1 @ y, loop1 @ x)
-c2t2c : (c1 : S1) (c2 : S1) -> IdP (<_> and S1 S1) (t2c (c2t' c1 c2)) (c1,c2) = split
+c2t2c : (c1 : S1) (c2 : S1) -> PathP (<_> and S1 S1) (t2c (c2t' c1 c2)) (c1,c2) = split
base -> c2t2c_base
-- again, I shouldn't need to do funext here;
-- I should be able to do a split inside of an interval binding
------------------------------------------------------------------------
-- combine everything to get that Torus = S1 * S1
-S1S1equalsTorus : Id U (and S1 S1) Torus = isoId (and S1 S1) Torus c2t t2c t2c2t rem
+S1S1equalsTorus : Path U (and S1 S1) Torus = isoPath (and S1 S1) Torus c2t t2c t2c2t rem
where
- rem (c:and S1 S1) : Id (and S1 S1) (t2c (c2t c)) c = c2t2c c.1 c.2
+ rem (c:and S1 S1) : Path (and S1 S1) (t2c (c2t c)) c = c2t2c c.1 c.2
-TorusEqualsS1S1 : Id U Torus (and S1 S1) = <i> S1S1equalsTorus @ -i
+TorusEqualsS1S1 : Path U Torus (and S1 S1) = <i> S1S1equalsTorus @ -i
-loopT : U = Id Torus ptT ptT
+loopT : U = Path Torus ptT ptT
--- funDep (A0 A1 :U) (p:Id U A0 A1) (u0:A0) (u1:A1) :
--- Id U (Id A0 u0 (transport (<i>p@-i) u1)) (Id A1 (transport p u0) u1) =
--- <i> Id (p @ i) (transport (<l> p @ (i/\l)) u0) (transport (<l> p @ (i\/-l)) u1)
+-- funDep (A0 A1 :U) (p:Path U A0 A1) (u0:A0) (u1:A1) :
+-- Path U (Path A0 u0 (transport (<i>p@-i) u1)) (Path A1 (transport p u0) u1) =
+-- <i> Path (p @ i) (transport (<l> p @ (i/\l)) u0) (transport (<l> p @ (i\/-l)) u1)
--- loopTorusEqualsZZ : Id U loopT (and Z Z) = <i> comp U (rem @ i) [(i = 1) -> rem1]
+-- loopTorusEqualsZZ : Path U loopT (and Z Z) = <i> comp U (rem @ i) [(i = 1) -> rem1]
-- where
--- rem : Id U loopT (Id (and S1 S1) (base,base) (base,base)) =
+-- rem : Path U loopT (Path (and S1 S1) (base,base) (base,base)) =
-- funDep Torus (and S1 S1) TorusEqualsS1S1 ptT (base,base)
--- rem1 : Id U (Id (and S1 S1) (base,base) (base,base)) (and Z Z) =
--- <i> comp U (lemIdAnd S1 S1 (base,base) (base,base) @ i)
+-- rem1 : Path U (Path (and S1 S1) (base,base) (base,base)) (and Z Z) =
+-- <i> comp U (lemPathAnd S1 S1 (base,base) (base,base) @ i)
-- [(i = 1) -> <j> and (loopS1equalsZ @ j) (loopS1equalsZ @ j)]
------------------------------------------------------------------------------
-- Proof of univalence using unglue:
-retIsContr (A B :U) (f:A->B) (g:B->A) (h : (x:A) -> Id A (g (f x)) x) (v:isContr B)
+retIsContr (A B :U) (f:A->B) (g:B->A) (h : (x:A) -> Path A (g (f x)) x) (v:isContr B)
: isContr A = (g b,p)
where
b : B = v.1
- q : (y:B) -> Id B b y = v.2
- p (x:A) : Id A (g b) x = <i>comp (<j>A) (g (q (f x)@i)) [(i=0) -> <j>g b,(i=1) -> h x]
+ q : (y:B) -> Path B b y = v.2
+ p (x:A) : Path A (g b) x = <i>comp (<j>A) (g (q (f x)@i)) [(i=0) -> <j>g b,(i=1) -> h x]
sigIsContr (A:U) (B:A->U) (u:isContr A) (q:(x:A) -> isContr (B x))
: isContr ((x:A)*B x) = ((a,g a),r)
where
a : A = u.1
- p : (x:A) -> Id A a x = u.2
+ p : (x:A) -> Path A a x = u.2
g (x:A) : B x = (q x).1
- h (x:A) : (y:B x) -> Id (B x) (g x) y = (q x).2
+ h (x:A) : (y:B x) -> Path (B x) (g x) y = (q x).2
C : U = (x:A) * B x
- r (z:C) : Id C (a,g a) z =
+ r (z:C) : Path C (a,g a) z =
<i>(p z.1@i,h (p z.1@i) (comp (<j>B (p z.1@i\/-j)) z.2 [(i=1)-><j>z.2])@i)
-isPathContr (A:U) (cA:isContr A) (x y:A) : isContr (Id A x y) = (p0,q)
+isPathContr (A:U) (cA:isContr A) (x y:A) : isContr (Path A x y) = (p0,q)
where
a : A = cA.1
- f : (x:A) -> Id A a x = cA.2
- p0 : Id A x y = <i>comp (<j>A) a [(i=0) -> f x,(i=1) -> f y]
- q (p:Id A x y) : Id (Id A x y) p0 p =
+ f : (x:A) -> Path A a x = cA.2
+ p0 : Path A x y = <i>comp (<j>A) a [(i=0) -> f x,(i=1) -> f y]
+ q (p:Path A x y) : Path (Path A x y) p0 p =
<j i>comp (<k>A) a [(i=0) -> f x,(i=1) -> f y,
(j=0) -> <k>comp (<l>A) a [(k=0) -> <l>a,(i=0) -> <l>f x@k/\l,(i=1) -> <l>f y@k/\l],
(j=1) -> f (p@i)]
isEquivContr (A B:U) (cA:isContr A) (cB:isContr B) (f:A->B) : isEquiv A B f =
- \ (y:B) -> sigIsContr A (\ (x:A) -> Id B y (f x)) cA (\ (x:A) -> isPathContr B cB y (f x))
+ \ (y:B) -> sigIsContr A (\ (x:A) -> Path B y (f x)) cA (\ (x:A) -> isPathContr B cB y (f x))
totalFun (A:U) (B C : A->U) (f : (x:A) -> B x -> C x) (w:Sigma A B) : Sigma A C =
(w.1,f (w.1) (w.2))
where
x : A = w.1.1
b : B x = w.1.2
- p : Id A x0 x = <i>(w.2@i).1
- q : IdP (<i>C (p@i)) z0 (f x b) = <i>(w.2@i).2
+ p : Path A x0 x = <i>(w.2@i).1
+ q : PathP (<i>C (p@i)) z0 (f x b) = <i>(w.2@i).2
b0 : B x0 = comp (<i>B (p@-i)) b []
- r : IdP (<i>B (p@-i)) b b0 = <i>comp (<j>B (p@-j\/-i)) b [(i=0) -> <k>b]
- s : Id (C x0) z0 (f x0 b0) = <i>comp (<j>C (p@(i/\-j))) (q@i) [(i=0) -> <k>z0,(i=1) -> <k>f (p@-k) (r@k)]
+ r : PathP (<i>B (p@-i)) b b0 = <i>comp (<j>B (p@-j\/-i)) b [(i=0) -> <k>b]
+ s : Path (C x0) z0 (f x0 b0) = <i>comp (<j>C (p@(i/\-j))) (q@i) [(i=0) -> <k>z0,(i=1) -> <k>f (p@-k) (r@k)]
compFunFib (A:U) (B C : A->U) (f : (x:A) -> B x -> C x) (x0:A) (z0:C x0) (u:fiber (B x0) (C x0) (f x0) z0) :
fiber (B x0) (C x0) (f x0) z0 = funFib2 A B C f x0 z0 (funFib1 A B C f x0 z0 u)
retFunFib (A:U) (B C : A->U) (f : (x:A) -> B x -> C x) (x0:A) (z0:C x0) (u:fiber (B x0) (C x0) (f x0) z0) :
- Id (fiber (B x0) (C x0) (f x0) z0) (funFib2 A B C f x0 z0 (funFib1 A B C f x0 z0 u)) u =
+ Path (fiber (B x0) (C x0) (f x0) z0) (funFib2 A B C f x0 z0 (funFib1 A B C f x0 z0 u)) u =
<l> (comp (<i> B x0) u.1 [(l=1) -> <j>u.1],
<i> comp (<j> C x0) (u.2 @ i) [ (l=1) -> <j> u.2 @ i,
(i = 0) -> <j> z0,
(isEquivContr (Sigma A B) (Sigma A C) cB cC (totalFun A B C f) (x,z))
-- test normal form
--- nequivFunFib : (A:U) (B C : A->U) (f : (x:A) -> B x -> C x) (cB : isContr (Sigma A B)) (cC : isContr (Sigma A C)) (x:A) ->isEquiv (B x) (C x) (f x) = \(A : U) -> \(B C : A -> U) -> \(f : (x : A) -> (B x) -> (C x)) -> \(cB : Sigma (Sigma A (\(x : A) -> B x)) (\(x : Sigma A (\(x : A) -> B x)) -> (y : Sigma A (\(x0 : A) -> B x0)) -> IdP (<!0> Sigma A (\(x0 : A) -> B x0)) x y)) -> \(cC : Sigma (Sigma A (\(x : A) -> C x)) (\(x : Sigma A (\(x : A) -> C x)) -> (y : Sigma A (\(x0 : A) -> C x0)) -> IdP (<!0> Sigma A (\(x0 : A) -> C x0)) x y)) -> \(x : A) -> \(z : C x) -> ((comp (<!0> B (comp (<!1> A) cC.1.1 [ (!0 = 0) -> <!1> (cC.2 ((cB.1.1,f cB.1.1 cB.1.2)) @ !1).1, (!0 = 1) -> <!1> (cC.2 ((x,z)) @ !1).1 ])) cB.1.2 [],<!0> comp (<!1> C (comp (<!2> A) cC.1.1 [ (!0 = 0) -> <!2> (cC.2 ((x,z)) @ !2).1, (!0 = 1)(!1 = 0) -> <!2> (cC.2 ((cB.1.1,f cB.1.1 cB.1.2)) @ !2).1, (!1 = 1) -> <!2> (cC.2 ((x,z)) @ !2).1 ])) (comp (<!1> C (comp (<!2> A) cC.1.1 [ (!0 = 0) -> <!2> (cC.2 ((x,z)) @ (!1 /\ !2)).1, (!0 = 1) -> <!2> (cC.2 ((cB.1.1,f cB.1.1 cB.1.2)) @ (!1 /\ !2)).1, (!1 = 0) -> <!2> cC.1.1 ])) cC.1.2 [ (!0 = 0) -> <!1> (cC.2 ((x,z)) @ !1).2, (!0 = 1) -> <!1> (cC.2 ((cB.1.1,f cB.1.1 cB.1.2)) @ !1).2 ]) [ (!0 = 0) -> <!1> z, (!0 = 1) -> <!1> f (comp (<!2> A) cC.1.1 [ (!1 = 0) -> <!2> (cC.2 ((cB.1.1,f cB.1.1 cB.1.2)) @ !2).1, (!1 = 1) -> <!2> (cC.2 ((x,z)) @ !2).1 ]) (comp (<!2> B (comp (<!3> A) cC.1.1 [ (!1 = 0) -> <!3> (cC.2 ((cB.1.1,f cB.1.1 cB.1.2)) @ !3).1, (!1 = 1)(!2 = 1) -> <!3> (cC.2 ((x,z)) @ !3).1, (!2 = 0) -> <!3> (cC.2 ((cB.1.1,f cB.1.1 cB.1.2)) @ !3).1 ])) cB.1.2 [ (!1 = 0) -> <!2> cB.1.2 ]) ]),\(x0 : Sigma (B x) (\(x0 : B x) -> IdP (<!0> C x) z (f x x0))) -> <!0> (comp (<!1> B x) (comp (<!1> B (comp (<!2> A) cC.1.1 [ (!0 = 0) -> <!2> comp (<!3> A) cC.1.1 [ (!1 = 0) -> <!3> (cC.2 ((cB.1.1,f cB.1.1 cB.1.2)) @ (!2 /\ !3)).1, (!1 = 1) -> <!3> (cC.2 ((x,z)) @ (!2 /\ !3)).1, (!2 = 0) -> <!3> cC.1.1 ], (!0 = 1) -> <!2> (cC.2 ((x,x0.2 @ -!1)) @ !2).1, (!1 = 0) -> <!2> (cC.2 (((cB.2 ((x,x0.1)) @ !0).1,f (cB.2 ((x,x0.1)) @ !0).1 (cB.2 ((x,x0.1)) @ !0).2)) @ !2).1, (!1 = 1) -> <!2> (cC.2 ((x,z)) @ !2).1 ])) (cB.2 ((x,x0.1)) @ !0).2 []) [ (!0 = 0) -> <!1> comp (<!0> B (comp (<!1> A) cC.1.1 [ (!0 = 0) -> <!1> (cC.2 ((cB.1.1,f cB.1.1 cB.1.2)) @ !1).1, (!0 = 1) -> <!1> (cC.2 ((x,z)) @ !1).1 ])) cB.1.2 [], (!0 = 1) -> <!1> comp (<!2> B x) x0.1 [ (!1 = 1) -> <!2> x0.1 ] ],<!2> comp (<!3> C x) (comp (<!3> C (comp (<!4> A) cC.1.1 [ (!0 = 0) -> <!4> comp (<!5> A) cC.1.1 [ (!2 = 0) -> <!5> (cC.2 ((x,z)) @ (!4 /\ !5)).1, (!2 = 1)(!3 = 0) -> <!5> (cC.2 ((cB.1.1,f cB.1.1 cB.1.2)) @ (!4 /\ !5)).1, (!3 = 1) -> <!5> (cC.2 ((x,z)) @ (!4 /\ !5)).1, (!4 = 0) -> <!5> cC.1.1 ], (!0 = 1) -> <!4> (cC.2 ((x,x0.2 @ (!2 /\ -!3))) @ !4).1, (!2 = 0) -> <!4> (cC.2 ((x,z)) @ !4).1, (!2 = 1)(!3 = 0) -> <!4> (cC.2 (((cB.2 ((x,x0.1)) @ !0).1,f (cB.2 ((x,x0.1)) @ !0).1 (cB.2 ((x,x0.1)) @ !0).2)) @ !4).1, (!3 = 1) -> <!4> (cC.2 ((x,z)) @ !4).1 ])) (comp (<!3> C (comp (<!4> A) cC.1.1 [ (!0 = 0) -> <!4> comp (<!5> A) cC.1.1 [ (!2 = 0) -> <!5> (cC.2 ((x,z)) @ (!3 /\ !4 /\ !5)).1, (!2 = 1) -> <!5> (cC.2 ((cB.1.1,f cB.1.1 cB.1.2)) @ (!3 /\ !4 /\ !5)).1, (!3 = 0) -> <!5> cC.1.1, (!4 = 0) -> <!5> cC.1.1 ], (!0 = 1) -> <!4> (cC.2 ((x,x0.2 @ !2)) @ (!3 /\ !4)).1, (!2 = 0) -> <!4> (cC.2 ((x,z)) @ (!3 /\ !4)).1, (!2 = 1) -> <!4> (cC.2 (((cB.2 ((x,x0.1)) @ !0).1,f (cB.2 ((x,x0.1)) @ !0).1 (cB.2 ((x,x0.1)) @ !0).2)) @ (!3 /\ !4)).1, (!3 = 0) -> <!4> cC.1.1 ])) cC.1.2 [ (!0 = 0) -> <!3> comp (<!4> C (comp (<!5> A) cC.1.1 [ (!2 = 0) -> <!5> (cC.2 ((x,z)) @ (!3 /\ !4 /\ !5)).1, (!2 = 1) -> <!5> (cC.2 ((cB.1.1,f cB.1.1 cB.1.2)) @ (!3 /\ !4 /\ !5)).1, (!3 = 0) -> <!5> cC.1.1, (!4 = 0) -> <!5> cC.1.1 ])) cC.1.2 [ (!2 = 0) -> <!4> (cC.2 ((x,z)) @ (!3 /\ !4)).2, (!2 = 1) -> <!4> (cC.2 ((cB.1.1,f cB.1.1 cB.1.2)) @ (!3 /\ !4)).2, (!3 = 0) -> <!4> cC.1.2 ], (!0 = 1) -> <!3> (cC.2 ((x,x0.2 @ !2)) @ !3).2, (!2 = 0) -> <!3> (cC.2 ((x,z)) @ !3).2, (!2 = 1) -> <!3> (cC.2 (((cB.2 ((x,x0.1)) @ !0).1,f (cB.2 ((x,x0.1)) @ !0).1 (cB.2 ((x,x0.1)) @ !0).2)) @ !3).2 ]) [ (!2 = 0) -> <!3> z, (!2 = 1) -> <!3> f (comp (<!4> A) cC.1.1 [ (!0 = 0) -> <!4> comp (<!5> A) cC.1.1 [ (!3 = 0) -> <!5> (cC.2 ((cB.1.1,f cB.1.1 cB.1.2)) @ (!4 /\ !5)).1, (!3 = 1) -> <!5> (cC.2 ((x,z)) @ (!4 /\ !5)).1, (!4 = 0) -> <!5> cC.1.1 ], (!0 = 1) -> <!4> (cC.2 ((x,x0.2 @ -!3)) @ !4).1, (!3 = 0) -> <!4> (cC.2 (((cB.2 ((x,x0.1)) @ !0).1,f (cB.2 ((x,x0.1)) @ !0).1 (cB.2 ((x,x0.1)) @ !0).2)) @ !4).1, (!3 = 1) -> <!4> (cC.2 ((x,z)) @ !4).1 ]) (comp (<!4> B (comp (<!5> A) cC.1.1 [ (!0 = 0) -> <!5> comp (<!6> A) cC.1.1 [ (!3 = 0) -> <!6> (cC.2 ((cB.1.1,f cB.1.1 cB.1.2)) @ (!5 /\ !6)).1, (!3 = 1)(!4 = 1) -> <!6> (cC.2 ((x,z)) @ (!5 /\ !6)).1, (!4 = 0) -> <!6> (cC.2 ((cB.1.1,f cB.1.1 cB.1.2)) @ (!5 /\ !6)).1, (!5 = 0) -> <!6> cC.1.1 ], (!0 = 1) -> <!5> (cC.2 ((x,x0.2 @ (-!3 \/ -!4))) @ !5).1, (!3 = 0) -> <!5> (cC.2 (((cB.2 ((x,x0.1)) @ !0).1,f (cB.2 ((x,x0.1)) @ !0).1 (cB.2 ((x,x0.1)) @ !0).2)) @ !5).1, (!3 = 1)(!4 = 1) -> <!5> (cC.2 ((x,z)) @ !5).1, (!4 = 0) -> <!5> (cC.2 (((cB.2 ((x,x0.1)) @ !0).1,f (cB.2 ((x,x0.1)) @ !0).1 (cB.2 ((x,x0.1)) @ !0).2)) @ !5).1 ])) (cB.2 ((x,x0.1)) @ !0).2 [ (!3 = 0) -> <!4> (cB.2 ((x,x0.1)) @ !0).2 ]) ]) [ (!0 = 0) -> <!3> comp (<!3> C (comp (<!4> A) cC.1.1 [ (!2 = 0) -> <!4> (cC.2 ((x,z)) @ !4).1, (!2 = 1)(!3 = 0) -> <!4> (cC.2 ((cB.1.1,f cB.1.1 cB.1.2)) @ !4).1, (!3 = 1) -> <!4> (cC.2 ((x,z)) @ !4).1 ])) (comp (<!3> C (comp (<!4> A) cC.1.1 [ (!2 = 0) -> <!4> (cC.2 ((x,z)) @ (!3 /\ !4)).1, (!2 = 1) -> <!4> (cC.2 ((cB.1.1,f cB.1.1 cB.1.2)) @ (!3 /\ !4)).1, (!3 = 0) -> <!4> cC.1.1 ])) cC.1.2 [ (!2 = 0) -> <!3> (cC.2 ((x,z)) @ !3).2, (!2 = 1) -> <!3> (cC.2 ((cB.1.1,f cB.1.1 cB.1.2)) @ !3).2 ]) [ (!2 = 0) -> <!3> z, (!2 = 1) -> <!3> f (comp (<!4> A) cC.1.1 [ (!3 = 0) -> <!4> (cC.2 ((cB.1.1,f cB.1.1 cB.1.2)) @ !4).1, (!3 = 1) -> <!4> (cC.2 ((x,z)) @ !4).1 ]) (comp (<!4> B (comp (<!5> A) cC.1.1 [ (!3 = 0) -> <!5> (cC.2 ((cB.1.1,f cB.1.1 cB.1.2)) @ !5).1, (!3 = 1)(!4 = 1) -> <!5> (cC.2 ((x,z)) @ !5).1, (!4 = 0) -> <!5> (cC.2 ((cB.1.1,f cB.1.1 cB.1.2)) @ !5).1 ])) cB.1.2 [ (!3 = 0) -> <!4> cB.1.2 ]) ], (!0 = 1) -> <!3> comp (<!4> C x) (x0.2 @ !2) [ (!2 = 0) -> <!4> z, (!2 = 1) -> <!4> f x (comp (<!5> B x) x0.1 [ (!3 = 1) -> <!5> x0.1, (!4 = 0) -> <!5> x0.1 ]), (!3 = 1) -> <!4> x0.2 @ !2 ], (!2 = 0) -> <!3> z, (!2 = 1) -> <!3> f x (comp (<!4> B x) (comp (<!1> B (comp (<!2> A) cC.1.1 [ (!0 = 0) -> <!2> comp (<!3> A) cC.1.1 [ (!1 = 0) -> <!3> (cC.2 ((cB.1.1,f cB.1.1 cB.1.2)) @ (!2 /\ !3)).1, (!1 = 1) -> <!3> (cC.2 ((x,z)) @ (!2 /\ !3)).1, (!2 = 0) -> <!3> cC.1.1 ], (!0 = 1) -> <!2> (cC.2 ((x,x0.2 @ -!1)) @ !2).1, (!1 = 0) -> <!2> (cC.2 (((cB.2 ((x,x0.1)) @ !0).1,f (cB.2 ((x,x0.1)) @ !0).1 (cB.2 ((x,x0.1)) @ !0).2)) @ !2).1, (!1 = 1) -> <!2> (cC.2 ((x,z)) @ !2).1 ])) (cB.2 ((x,x0.1)) @ !0).2 []) [ (!0 = 0) -> <!4> comp (<!0> B (comp (<!1> A) cC.1.1 [ (!0 = 0) -> <!1> (cC.2 ((cB.1.1,f cB.1.1 cB.1.2)) @ !1).1, (!0 = 1) -> <!1> (cC.2 ((x,z)) @ !1).1 ])) cB.1.2 [], (!0 = 1) -> <!4> comp (<!5> B x) x0.1 [ (!3 = 1)(!4 = 1) -> <!5> x0.1 ], (!3 = 0) -> <!4> comp (<!1> B (comp (<!2> A) cC.1.1 [ (!0 = 0) -> <!2> comp (<!3> A) cC.1.1 [ (!1 = 0) -> <!3> (cC.2 ((cB.1.1,f cB.1.1 cB.1.2)) @ (!2 /\ !3)).1, (!1 = 1) -> <!3> (cC.2 ((x,z)) @ (!2 /\ !3)).1, (!2 = 0) -> <!3> cC.1.1 ], (!0 = 1) -> <!2> (cC.2 ((x,x0.2 @ -!1)) @ !2).1, (!1 = 0) -> <!2> (cC.2 (((cB.2 ((x,x0.1)) @ !0).1,f (cB.2 ((x,x0.1)) @ !0).1 (cB.2 ((x,x0.1)) @ !0).2)) @ !2).1, (!1 = 1) -> <!2> (cC.2 ((x,z)) @ !2).1 ])) (cB.2 ((x,x0.1)) @ !0).2 [] ]) ]))
+-- nequivFunFib : (A:U) (B C : A->U) (f : (x:A) -> B x -> C x) (cB : isContr (Sigma A B)) (cC : isContr (Sigma A C)) (x:A) ->isEquiv (B x) (C x) (f x) = \(A : U) -> \(B C : A -> U) -> \(f : (x : A) -> (B x) -> (C x)) -> \(cB : Sigma (Sigma A (\(x : A) -> B x)) (\(x : Sigma A (\(x : A) -> B x)) -> (y : Sigma A (\(x0 : A) -> B x0)) -> PathP (<!0> Sigma A (\(x0 : A) -> B x0)) x y)) -> \(cC : Sigma (Sigma A (\(x : A) -> C x)) (\(x : Sigma A (\(x : A) -> C x)) -> (y : Sigma A (\(x0 : A) -> C x0)) -> PathP (<!0> Sigma A (\(x0 : A) -> C x0)) x y)) -> \(x : A) -> \(z : C x) -> ((comp (<!0> B (comp (<!1> A) cC.1.1 [ (!0 = 0) -> <!1> (cC.2 ((cB.1.1,f cB.1.1 cB.1.2)) @ !1).1, (!0 = 1) -> <!1> (cC.2 ((x,z)) @ !1).1 ])) cB.1.2 [],<!0> comp (<!1> C (comp (<!2> A) cC.1.1 [ (!0 = 0) -> <!2> (cC.2 ((x,z)) @ !2).1, (!0 = 1)(!1 = 0) -> <!2> (cC.2 ((cB.1.1,f cB.1.1 cB.1.2)) @ !2).1, (!1 = 1) -> <!2> (cC.2 ((x,z)) @ !2).1 ])) (comp (<!1> C (comp (<!2> A) cC.1.1 [ (!0 = 0) -> <!2> (cC.2 ((x,z)) @ (!1 /\ !2)).1, (!0 = 1) -> <!2> (cC.2 ((cB.1.1,f cB.1.1 cB.1.2)) @ (!1 /\ !2)).1, (!1 = 0) -> <!2> cC.1.1 ])) cC.1.2 [ (!0 = 0) -> <!1> (cC.2 ((x,z)) @ !1).2, (!0 = 1) -> <!1> (cC.2 ((cB.1.1,f cB.1.1 cB.1.2)) @ !1).2 ]) [ (!0 = 0) -> <!1> z, (!0 = 1) -> <!1> f (comp (<!2> A) cC.1.1 [ (!1 = 0) -> <!2> (cC.2 ((cB.1.1,f cB.1.1 cB.1.2)) @ !2).1, (!1 = 1) -> <!2> (cC.2 ((x,z)) @ !2).1 ]) (comp (<!2> B (comp (<!3> A) cC.1.1 [ (!1 = 0) -> <!3> (cC.2 ((cB.1.1,f cB.1.1 cB.1.2)) @ !3).1, (!1 = 1)(!2 = 1) -> <!3> (cC.2 ((x,z)) @ !3).1, (!2 = 0) -> <!3> (cC.2 ((cB.1.1,f cB.1.1 cB.1.2)) @ !3).1 ])) cB.1.2 [ (!1 = 0) -> <!2> cB.1.2 ]) ]),\(x0 : Sigma (B x) (\(x0 : B x) -> PathP (<!0> C x) z (f x x0))) -> <!0> (comp (<!1> B x) (comp (<!1> B (comp (<!2> A) cC.1.1 [ (!0 = 0) -> <!2> comp (<!3> A) cC.1.1 [ (!1 = 0) -> <!3> (cC.2 ((cB.1.1,f cB.1.1 cB.1.2)) @ (!2 /\ !3)).1, (!1 = 1) -> <!3> (cC.2 ((x,z)) @ (!2 /\ !3)).1, (!2 = 0) -> <!3> cC.1.1 ], (!0 = 1) -> <!2> (cC.2 ((x,x0.2 @ -!1)) @ !2).1, (!1 = 0) -> <!2> (cC.2 (((cB.2 ((x,x0.1)) @ !0).1,f (cB.2 ((x,x0.1)) @ !0).1 (cB.2 ((x,x0.1)) @ !0).2)) @ !2).1, (!1 = 1) -> <!2> (cC.2 ((x,z)) @ !2).1 ])) (cB.2 ((x,x0.1)) @ !0).2 []) [ (!0 = 0) -> <!1> comp (<!0> B (comp (<!1> A) cC.1.1 [ (!0 = 0) -> <!1> (cC.2 ((cB.1.1,f cB.1.1 cB.1.2)) @ !1).1, (!0 = 1) -> <!1> (cC.2 ((x,z)) @ !1).1 ])) cB.1.2 [], (!0 = 1) -> <!1> comp (<!2> B x) x0.1 [ (!1 = 1) -> <!2> x0.1 ] ],<!2> comp (<!3> C x) (comp (<!3> C (comp (<!4> A) cC.1.1 [ (!0 = 0) -> <!4> comp (<!5> A) cC.1.1 [ (!2 = 0) -> <!5> (cC.2 ((x,z)) @ (!4 /\ !5)).1, (!2 = 1)(!3 = 0) -> <!5> (cC.2 ((cB.1.1,f cB.1.1 cB.1.2)) @ (!4 /\ !5)).1, (!3 = 1) -> <!5> (cC.2 ((x,z)) @ (!4 /\ !5)).1, (!4 = 0) -> <!5> cC.1.1 ], (!0 = 1) -> <!4> (cC.2 ((x,x0.2 @ (!2 /\ -!3))) @ !4).1, (!2 = 0) -> <!4> (cC.2 ((x,z)) @ !4).1, (!2 = 1)(!3 = 0) -> <!4> (cC.2 (((cB.2 ((x,x0.1)) @ !0).1,f (cB.2 ((x,x0.1)) @ !0).1 (cB.2 ((x,x0.1)) @ !0).2)) @ !4).1, (!3 = 1) -> <!4> (cC.2 ((x,z)) @ !4).1 ])) (comp (<!3> C (comp (<!4> A) cC.1.1 [ (!0 = 0) -> <!4> comp (<!5> A) cC.1.1 [ (!2 = 0) -> <!5> (cC.2 ((x,z)) @ (!3 /\ !4 /\ !5)).1, (!2 = 1) -> <!5> (cC.2 ((cB.1.1,f cB.1.1 cB.1.2)) @ (!3 /\ !4 /\ !5)).1, (!3 = 0) -> <!5> cC.1.1, (!4 = 0) -> <!5> cC.1.1 ], (!0 = 1) -> <!4> (cC.2 ((x,x0.2 @ !2)) @ (!3 /\ !4)).1, (!2 = 0) -> <!4> (cC.2 ((x,z)) @ (!3 /\ !4)).1, (!2 = 1) -> <!4> (cC.2 (((cB.2 ((x,x0.1)) @ !0).1,f (cB.2 ((x,x0.1)) @ !0).1 (cB.2 ((x,x0.1)) @ !0).2)) @ (!3 /\ !4)).1, (!3 = 0) -> <!4> cC.1.1 ])) cC.1.2 [ (!0 = 0) -> <!3> comp (<!4> C (comp (<!5> A) cC.1.1 [ (!2 = 0) -> <!5> (cC.2 ((x,z)) @ (!3 /\ !4 /\ !5)).1, (!2 = 1) -> <!5> (cC.2 ((cB.1.1,f cB.1.1 cB.1.2)) @ (!3 /\ !4 /\ !5)).1, (!3 = 0) -> <!5> cC.1.1, (!4 = 0) -> <!5> cC.1.1 ])) cC.1.2 [ (!2 = 0) -> <!4> (cC.2 ((x,z)) @ (!3 /\ !4)).2, (!2 = 1) -> <!4> (cC.2 ((cB.1.1,f cB.1.1 cB.1.2)) @ (!3 /\ !4)).2, (!3 = 0) -> <!4> cC.1.2 ], (!0 = 1) -> <!3> (cC.2 ((x,x0.2 @ !2)) @ !3).2, (!2 = 0) -> <!3> (cC.2 ((x,z)) @ !3).2, (!2 = 1) -> <!3> (cC.2 (((cB.2 ((x,x0.1)) @ !0).1,f (cB.2 ((x,x0.1)) @ !0).1 (cB.2 ((x,x0.1)) @ !0).2)) @ !3).2 ]) [ (!2 = 0) -> <!3> z, (!2 = 1) -> <!3> f (comp (<!4> A) cC.1.1 [ (!0 = 0) -> <!4> comp (<!5> A) cC.1.1 [ (!3 = 0) -> <!5> (cC.2 ((cB.1.1,f cB.1.1 cB.1.2)) @ (!4 /\ !5)).1, (!3 = 1) -> <!5> (cC.2 ((x,z)) @ (!4 /\ !5)).1, (!4 = 0) -> <!5> cC.1.1 ], (!0 = 1) -> <!4> (cC.2 ((x,x0.2 @ -!3)) @ !4).1, (!3 = 0) -> <!4> (cC.2 (((cB.2 ((x,x0.1)) @ !0).1,f (cB.2 ((x,x0.1)) @ !0).1 (cB.2 ((x,x0.1)) @ !0).2)) @ !4).1, (!3 = 1) -> <!4> (cC.2 ((x,z)) @ !4).1 ]) (comp (<!4> B (comp (<!5> A) cC.1.1 [ (!0 = 0) -> <!5> comp (<!6> A) cC.1.1 [ (!3 = 0) -> <!6> (cC.2 ((cB.1.1,f cB.1.1 cB.1.2)) @ (!5 /\ !6)).1, (!3 = 1)(!4 = 1) -> <!6> (cC.2 ((x,z)) @ (!5 /\ !6)).1, (!4 = 0) -> <!6> (cC.2 ((cB.1.1,f cB.1.1 cB.1.2)) @ (!5 /\ !6)).1, (!5 = 0) -> <!6> cC.1.1 ], (!0 = 1) -> <!5> (cC.2 ((x,x0.2 @ (-!3 \/ -!4))) @ !5).1, (!3 = 0) -> <!5> (cC.2 (((cB.2 ((x,x0.1)) @ !0).1,f (cB.2 ((x,x0.1)) @ !0).1 (cB.2 ((x,x0.1)) @ !0).2)) @ !5).1, (!3 = 1)(!4 = 1) -> <!5> (cC.2 ((x,z)) @ !5).1, (!4 = 0) -> <!5> (cC.2 (((cB.2 ((x,x0.1)) @ !0).1,f (cB.2 ((x,x0.1)) @ !0).1 (cB.2 ((x,x0.1)) @ !0).2)) @ !5).1 ])) (cB.2 ((x,x0.1)) @ !0).2 [ (!3 = 0) -> <!4> (cB.2 ((x,x0.1)) @ !0).2 ]) ]) [ (!0 = 0) -> <!3> comp (<!3> C (comp (<!4> A) cC.1.1 [ (!2 = 0) -> <!4> (cC.2 ((x,z)) @ !4).1, (!2 = 1)(!3 = 0) -> <!4> (cC.2 ((cB.1.1,f cB.1.1 cB.1.2)) @ !4).1, (!3 = 1) -> <!4> (cC.2 ((x,z)) @ !4).1 ])) (comp (<!3> C (comp (<!4> A) cC.1.1 [ (!2 = 0) -> <!4> (cC.2 ((x,z)) @ (!3 /\ !4)).1, (!2 = 1) -> <!4> (cC.2 ((cB.1.1,f cB.1.1 cB.1.2)) @ (!3 /\ !4)).1, (!3 = 0) -> <!4> cC.1.1 ])) cC.1.2 [ (!2 = 0) -> <!3> (cC.2 ((x,z)) @ !3).2, (!2 = 1) -> <!3> (cC.2 ((cB.1.1,f cB.1.1 cB.1.2)) @ !3).2 ]) [ (!2 = 0) -> <!3> z, (!2 = 1) -> <!3> f (comp (<!4> A) cC.1.1 [ (!3 = 0) -> <!4> (cC.2 ((cB.1.1,f cB.1.1 cB.1.2)) @ !4).1, (!3 = 1) -> <!4> (cC.2 ((x,z)) @ !4).1 ]) (comp (<!4> B (comp (<!5> A) cC.1.1 [ (!3 = 0) -> <!5> (cC.2 ((cB.1.1,f cB.1.1 cB.1.2)) @ !5).1, (!3 = 1)(!4 = 1) -> <!5> (cC.2 ((x,z)) @ !5).1, (!4 = 0) -> <!5> (cC.2 ((cB.1.1,f cB.1.1 cB.1.2)) @ !5).1 ])) cB.1.2 [ (!3 = 0) -> <!4> cB.1.2 ]) ], (!0 = 1) -> <!3> comp (<!4> C x) (x0.2 @ !2) [ (!2 = 0) -> <!4> z, (!2 = 1) -> <!4> f x (comp (<!5> B x) x0.1 [ (!3 = 1) -> <!5> x0.1, (!4 = 0) -> <!5> x0.1 ]), (!3 = 1) -> <!4> x0.2 @ !2 ], (!2 = 0) -> <!3> z, (!2 = 1) -> <!3> f x (comp (<!4> B x) (comp (<!1> B (comp (<!2> A) cC.1.1 [ (!0 = 0) -> <!2> comp (<!3> A) cC.1.1 [ (!1 = 0) -> <!3> (cC.2 ((cB.1.1,f cB.1.1 cB.1.2)) @ (!2 /\ !3)).1, (!1 = 1) -> <!3> (cC.2 ((x,z)) @ (!2 /\ !3)).1, (!2 = 0) -> <!3> cC.1.1 ], (!0 = 1) -> <!2> (cC.2 ((x,x0.2 @ -!1)) @ !2).1, (!1 = 0) -> <!2> (cC.2 (((cB.2 ((x,x0.1)) @ !0).1,f (cB.2 ((x,x0.1)) @ !0).1 (cB.2 ((x,x0.1)) @ !0).2)) @ !2).1, (!1 = 1) -> <!2> (cC.2 ((x,z)) @ !2).1 ])) (cB.2 ((x,x0.1)) @ !0).2 []) [ (!0 = 0) -> <!4> comp (<!0> B (comp (<!1> A) cC.1.1 [ (!0 = 0) -> <!1> (cC.2 ((cB.1.1,f cB.1.1 cB.1.2)) @ !1).1, (!0 = 1) -> <!1> (cC.2 ((x,z)) @ !1).1 ])) cB.1.2 [], (!0 = 1) -> <!4> comp (<!5> B x) x0.1 [ (!3 = 1)(!4 = 1) -> <!5> x0.1 ], (!3 = 0) -> <!4> comp (<!1> B (comp (<!2> A) cC.1.1 [ (!0 = 0) -> <!2> comp (<!3> A) cC.1.1 [ (!1 = 0) -> <!3> (cC.2 ((cB.1.1,f cB.1.1 cB.1.2)) @ (!2 /\ !3)).1, (!1 = 1) -> <!3> (cC.2 ((x,z)) @ (!2 /\ !3)).1, (!2 = 0) -> <!3> cC.1.1 ], (!0 = 1) -> <!2> (cC.2 ((x,x0.2 @ -!1)) @ !2).1, (!1 = 0) -> <!2> (cC.2 (((cB.2 ((x,x0.1)) @ !0).1,f (cB.2 ((x,x0.1)) @ !0).1 (cB.2 ((x,x0.1)) @ !0).2)) @ !2).1, (!1 = 1) -> <!2> (cC.2 ((x,z)) @ !2).1 ])) (cB.2 ((x,x0.1)) @ !0).2 [] ]) ]))
lem1 (B:U) : isContr ((X:U) * equiv X B) =
,fill (<j> B) b [(i=0) -> <j> b
,(i=1) -> (w.2.2 b).1.2])
contr (v : fiber glueB B unglueElemB b)
- : Id (fiber glueB B unglueElemB b) center v
+ : Path (fiber glueB B unglueElemB b) center v
= <j> (glueElem (comp (<j> B) b [(i=0) -> <k> v.2 @ (j /\ k)
,(i=1) -> ((w.2.2 b).2 v @ j).2
,(j=0) -> fill (<j> B) b [(i=0) -> <j> b
,(j=1) -> v.2])
in (center,contr)))
--- nlem1 (B:U) : isContr ((X:U) * equiv X B) = ((B,(\(x : B) -> x,\(a : B) -> ((a,<!0> a),\(z : Sigma B (\(x : B) -> IdP (<!0> B) a x)) -> <!0> (z.2 @ !0,<!1> z.2 @ (!0 /\ !1))))),\(w : Sigma U (\(X : U) -> Sigma (X -> B) (\(f : X -> B) -> (y : B) -> Sigma (Sigma X (\(x : X) -> IdP (<!0> B) y (f x))) (\(x : Sigma X (\(x : X) -> IdP (<!0> B) y (f x))) -> (y0 : Sigma X (\(x0 : X) -> IdP (<!0> B) y (f x0))) -> IdP (<!0> Sigma X (\(x0 : X) -> IdP (<!0> B) y (f x0))) x y0)))) -> <!0> (glue B [ (!0 = 0) -> (B,(\(x : B) -> x,\(a : B) -> ((a,<!0> a),\(z : Sigma B (\(x : B) -> IdP (<!0> B) a x)) -> <!0> (z.2 @ !0,<!1> z.2 @ (!0 /\ !1))))), (!0 = 1) -> w ],(\(g : glue B [ (!0 = 0) -> (B,(\(x : B) -> x,\(a : B) -> ((a,<!0> a),\(z : Sigma B (\(x : B) -> IdP (<!0> B) a x)) -> <!0> (z.2 @ !0,<!1> z.2 @ (!0 /\ !1))))), (!0 = 1) -> w ]) -> unglueElem g [ (!0 = 0) -> (B,(\(x : B) -> x,\(a : B) -> ((a,<!0> a),\(z : Sigma B (\(x : B) -> IdP (<!0> B) a x)) -> <!0> (z.2 @ !0,<!1> z.2 @ (!0 /\ !1))))), (!0 = 1) -> w ],\(b : B) -> ((glueElem (comp (<!1> B) b [ (!0 = 0) -> <!1> b, (!0 = 1) -> <!1> (w.2.2 b).1.2 @ !1 ]) [ (!0 = 0) -> b, (!0 = 1) -> (w.2.2 b).1.1 ],<!1> comp (<!2> B) b [ (!0 = 0) -> <!2> b, (!0 = 1) -> <!2> (w.2.2 b).1.2 @ (!1 /\ !2), (!1 = 0) -> <!2> b ]),\(v : Sigma (glue B [ (!0 = 0) -> (B,(\(x : B) -> x,\(a : B) -> ((a,<!0> a),\(z : Sigma B (\(x : B) -> IdP (<!0> B) a x)) -> <!0> (z.2 @ !0,<!1> z.2 @ (!0 /\ !1))))), (!0 = 1) -> w ]) (\(x : glue B [ (!0 = 0) -> (B,(\(x : B) -> x,\(a : B) -> ((a,<!0> a),\(z : Sigma B (\(x : B) -> IdP (<!0> B) a x)) -> <!0> (z.2 @ !0,<!1> z.2 @ (!0 /\ !1))))), (!0 = 1) -> w ]) -> IdP (<!1> B) b (unglueElem x [ (!0 = 0) -> (B,(\(x0 : B) -> x0,\(a : B) -> ((a,<!0> a),\(z : Sigma B (\(x0 : B) -> IdP (<!0> B) a x0)) -> <!0> (z.2 @ !0,<!1> z.2 @ (!0 /\ !1))))), (!0 = 1) -> w ]))) -> <!1> (glueElem (comp (<!2> B) b [ (!0 = 0) -> <!2> v.2 @ (!1 /\ !2), (!0 = 1) -> <!2> ((w.2.2 b).2 v @ !1).2 @ !2, (!1 = 0) -> <!2> comp (<!3> B) b [ (!0 = 0) -> <!3> b, (!0 = 1) -> <!3> (w.2.2 b).1.2 @ (!2 /\ !3), (!2 = 0) -> <!3> b ], (!1 = 1) -> <!2> v.2 @ !2 ]) [ (!0 = 0) -> v.2 @ !1, (!0 = 1) -> ((w.2.2 b).2 v @ !1).1 ],<!2> comp (<!3> B) b [ (!0 = 0) -> <!3> v.2 @ (!1 /\ !2 /\ !3), (!0 = 1) -> <!3> ((w.2.2 b).2 v @ !1).2 @ (!2 /\ !3), (!1 = 0) -> <!3> comp (<!4> B) b [ (!0 = 0) -> <!4> b, (!0 = 1) -> <!4> (w.2.2 b).1.2 @ (!2 /\ !3 /\ !4), (!2 = 0) -> <!4> b, (!3 = 0) -> <!4> b ], (!1 = 1) -> <!3> v.2 @ (!2 /\ !3), (!2 = 0) -> <!3> b ])))))
+-- nlem1 (B:U) : isContr ((X:U) * equiv X B) = ((B,(\(x : B) -> x,\(a : B) -> ((a,<!0> a),\(z : Sigma B (\(x : B) -> PathP (<!0> B) a x)) -> <!0> (z.2 @ !0,<!1> z.2 @ (!0 /\ !1))))),\(w : Sigma U (\(X : U) -> Sigma (X -> B) (\(f : X -> B) -> (y : B) -> Sigma (Sigma X (\(x : X) -> PathP (<!0> B) y (f x))) (\(x : Sigma X (\(x : X) -> PathP (<!0> B) y (f x))) -> (y0 : Sigma X (\(x0 : X) -> PathP (<!0> B) y (f x0))) -> PathP (<!0> Sigma X (\(x0 : X) -> PathP (<!0> B) y (f x0))) x y0)))) -> <!0> (glue B [ (!0 = 0) -> (B,(\(x : B) -> x,\(a : B) -> ((a,<!0> a),\(z : Sigma B (\(x : B) -> PathP (<!0> B) a x)) -> <!0> (z.2 @ !0,<!1> z.2 @ (!0 /\ !1))))), (!0 = 1) -> w ],(\(g : glue B [ (!0 = 0) -> (B,(\(x : B) -> x,\(a : B) -> ((a,<!0> a),\(z : Sigma B (\(x : B) -> PathP (<!0> B) a x)) -> <!0> (z.2 @ !0,<!1> z.2 @ (!0 /\ !1))))), (!0 = 1) -> w ]) -> unglueElem g [ (!0 = 0) -> (B,(\(x : B) -> x,\(a : B) -> ((a,<!0> a),\(z : Sigma B (\(x : B) -> PathP (<!0> B) a x)) -> <!0> (z.2 @ !0,<!1> z.2 @ (!0 /\ !1))))), (!0 = 1) -> w ],\(b : B) -> ((glueElem (comp (<!1> B) b [ (!0 = 0) -> <!1> b, (!0 = 1) -> <!1> (w.2.2 b).1.2 @ !1 ]) [ (!0 = 0) -> b, (!0 = 1) -> (w.2.2 b).1.1 ],<!1> comp (<!2> B) b [ (!0 = 0) -> <!2> b, (!0 = 1) -> <!2> (w.2.2 b).1.2 @ (!1 /\ !2), (!1 = 0) -> <!2> b ]),\(v : Sigma (glue B [ (!0 = 0) -> (B,(\(x : B) -> x,\(a : B) -> ((a,<!0> a),\(z : Sigma B (\(x : B) -> PathP (<!0> B) a x)) -> <!0> (z.2 @ !0,<!1> z.2 @ (!0 /\ !1))))), (!0 = 1) -> w ]) (\(x : glue B [ (!0 = 0) -> (B,(\(x : B) -> x,\(a : B) -> ((a,<!0> a),\(z : Sigma B (\(x : B) -> PathP (<!0> B) a x)) -> <!0> (z.2 @ !0,<!1> z.2 @ (!0 /\ !1))))), (!0 = 1) -> w ]) -> PathP (<!1> B) b (unglueElem x [ (!0 = 0) -> (B,(\(x0 : B) -> x0,\(a : B) -> ((a,<!0> a),\(z : Sigma B (\(x0 : B) -> PathP (<!0> B) a x0)) -> <!0> (z.2 @ !0,<!1> z.2 @ (!0 /\ !1))))), (!0 = 1) -> w ]))) -> <!1> (glueElem (comp (<!2> B) b [ (!0 = 0) -> <!2> v.2 @ (!1 /\ !2), (!0 = 1) -> <!2> ((w.2.2 b).2 v @ !1).2 @ !2, (!1 = 0) -> <!2> comp (<!3> B) b [ (!0 = 0) -> <!3> b, (!0 = 1) -> <!3> (w.2.2 b).1.2 @ (!2 /\ !3), (!2 = 0) -> <!3> b ], (!1 = 1) -> <!2> v.2 @ !2 ]) [ (!0 = 0) -> v.2 @ !1, (!0 = 1) -> ((w.2.2 b).2 v @ !1).1 ],<!2> comp (<!3> B) b [ (!0 = 0) -> <!3> v.2 @ (!1 /\ !2 /\ !3), (!0 = 1) -> <!3> ((w.2.2 b).2 v @ !1).2 @ (!2 /\ !3), (!1 = 0) -> <!3> comp (<!4> B) b [ (!0 = 0) -> <!4> b, (!0 = 1) -> <!4> (w.2.2 b).1.2 @ (!2 /\ !3 /\ !4), (!2 = 0) -> <!4> b, (!3 = 0) -> <!4> b ], (!1 = 1) -> <!3> v.2 @ (!2 /\ !3), (!2 = 0) -> <!3> b ])))))
-contrSingl' (A : U) (a b : A) (p : Id A a b) :
- Id ((x:A) * Id A x b) (b,refl A b) (a,p) = <i> (p @ -i,<j> p @ -i\/j)
+contrSingl' (A : U) (a b : A) (p : Path A a b) :
+ Path ((x:A) * Path A x b) (b,refl A b) (a,p) = <i> (p @ -i,<j> p @ -i\/j)
-lemSinglContr' (A:U) (a:A) : isContr ((x:A) * Id A x a) =
- ((a,refl A a),\ (z:(x:A) * Id A x a) -> contrSingl' A z.1 a z.2)
+lemSinglContr' (A:U) (a:A) : isContr ((x:A) * Path A x a) =
+ ((a,refl A a),\ (z:(x:A) * Path A x a) -> contrSingl' A z.1 a z.2)
-thmUniv (t : (A X : U) -> Id U X A -> equiv X A) (A : U) :
- (X : U) -> isEquiv (Id U X A) (equiv X A) (t A X) =
- equivFunFib U (\(X : U) -> Id U X A) (\(X : U) -> equiv X A)
+thmUniv (t : (A X : U) -> Path U X A -> equiv X A) (A : U) :
+ (X : U) -> isEquiv (Path U X A) (equiv X A) (t A X) =
+ equivFunFib U (\(X : U) -> Path U X A) (\(X : U) -> equiv X A)
(t A) (lemSinglContr' U A) (lem1 A)
-transEquiv' (A X : U) (p : Id U X A) : equiv X A =
+transEquiv' (A X : U) (p : Path U X A) : equiv X A =
substTrans U (\(Y : U) -> equiv Y A) A X (<i> p @ -i) (idEquiv A)
-- The univalence axiom
-univalence (A X : U) : isEquiv (Id U X A) (equiv X A) (transEquiv' A X) =
+univalence (A X : U) : isEquiv (Path U X A) (equiv X A) (transEquiv' A X) =
thmUniv transEquiv' A X
-- The standard formulation of univalence whose normal can be computed:
-corrUniv (A B : U) : Id U (Id U A B) (equiv A B) =
- equivId (Id U A B) (equiv A B) (transEquiv' B A) (univalence B A)
+corrUniv (A B : U) : Path U (Path U A B) (equiv A B) =
+ equivPath (Path U A B) (equiv A B) (transEquiv' B A) (univalence B A)
-corrUniv' (A B : U) : equiv (Id U A B) (equiv A B) =
+corrUniv' (A B : U) : equiv (Path U A B) (equiv A B) =
(transEquiv' B A,univalence B A)
-- transEquiv is an equiv
transEquivIsEquiv (A B : U)
- : isEquiv (Id U A B) (equiv A B) (transEquiv A B)
- = gradLemma (Id U A B) (equiv A B) (transEquiv A B)
- (transEquivToId A B) (idToId A B) (eqToEq A B)
+ : isEquiv (Path U A B) (equiv A B) (transEquiv A B)
+ = gradLemma (Path U A B) (equiv A B) (transEquiv A B)
+ (transEquivToPath A B) (idToPath A B) (eqToEq A B)
-- Univalence proved using transEquiv
-- univalenceTrans takes extremely much memory when normalizing
-univalenceTrans (A B:U) : Id U (Id U A B) (equiv A B) =
- isoId (Id U A B) (equiv A B) (transEquiv A B)
- (transEquivToId A B) (idToId A B) (eqToEq A B)
+univalenceTrans (A B:U) : Path U (Path U A B) (equiv A B) =
+ isoPath (Path U A B) (equiv A B) (transEquiv A B)
+ (transEquivToPath A B) (idToPath A B) (eqToEq A B)
-univalenceTrans' (A B : U) : equiv (Id U A B) (equiv A B) =
+univalenceTrans' (A B : U) : equiv (Path U A B) (equiv A B) =
(transEquiv A B,transEquivIsEquiv A B)
-- This also takes too long to normalize:
-slowtest (A : U) : Id (equiv A A)
- (transEquiv A A (transEquivToId A A (idEquiv A))) (idEquiv A) =
- idToId A A (idEquiv A)
+slowtest (A : U) : Path (equiv A A)
+ (transEquiv A A (transEquivToPath A A (idEquiv A))) (idEquiv A) =
+ idToPath A A (idEquiv A)
-- ------------------------------------------------------------------------------
-- -- TODO: Adapt this to new definition of equiv
-- -- The canonical map defined using J
--- -- canIdToEquiv (A : U) : (B : U) -> Id U A B -> equiv A B =
--- -- J U A (\ (B : U) (_ : Id U A B) -> equiv A B) (idEquiv A)
+-- -- canPathToEquiv (A : U) : (B : U) -> Path U A B -> equiv A B =
+-- -- J U A (\ (B : U) (_ : Path U A B) -> equiv A B) (idEquiv A)
--- -- canDiagTrans (A : U) : Id (A -> A) (canIdToEquiv A A (<_> A)).1 (idfun A) =
+-- -- canDiagTrans (A : U) : Path (A -> A) (canPathToEquiv A A (<_> A)).1 (idfun A) =
-- -- <i> fill (<_> A -> A) (idfun A) [] @ -i
--- -- transDiagTrans (A : U) : Id (A -> A) (idfun A) (trans A A (<_> A)) =
+-- -- transDiagTrans (A : U) : Path (A -> A) (idfun A) (trans A A (<_> A)) =
-- -- <i> \ (a : A) -> fill (<_> A) a [] @ i
--- -- canIdToEquivLem (A : U) : (B : U) (p : Id U A B) ->
--- -- Id (A -> B) (canIdToEquiv A B p).1 (transEquiv A B p).1
+-- -- canPathToEquivLem (A : U) : (B : U) (p : Path U A B) ->
+-- -- Path (A -> B) (canPathToEquiv A B p).1 (transEquiv A B p).1
-- -- = J U A
--- -- (\ (B : U) (p : Id U A B) ->
--- -- Id (A -> B) (canIdToEquiv A B p).1 (transEquiv A B p).1)
--- -- (compId (A -> A)
--- -- (canIdToEquiv A A (<_> A)).1 (idfun A) (trans A A (<_> A))
+-- -- (\ (B : U) (p : Path U A B) ->
+-- -- Path (A -> B) (canPathToEquiv A B p).1 (transEquiv A B p).1)
+-- -- (compPath (A -> A)
+-- -- (canPathToEquiv A A (<_> A)).1 (idfun A) (trans A A (<_> A))
-- -- (canDiagTrans A) (transDiagTrans A))
--- -- canIdToEquivIsTransEquiv (A B : U) :
--- -- Id (Id U A B -> equiv A B) (canIdToEquiv A B) (transEquiv A B) =
--- -- <i> \ (p : Id U A B) ->
--- -- equivLemma A B (canIdToEquiv A B p) (transEquiv A B p)
--- -- (canIdToEquivLem A B p) @ i
+-- -- canPathToEquivIsTransEquiv (A B : U) :
+-- -- Path (Path U A B -> equiv A B) (canPathToEquiv A B) (transEquiv A B) =
+-- -- <i> \ (p : Path U A B) ->
+-- -- equivLemma A B (canPathToEquiv A B p) (transEquiv A B p)
+-- -- (canPathToEquivLem A B p) @ i
-- -- -- The canonical map is an equivalence
--- -- univalenceJ (A B : U) : isEquiv (Id U A B) (equiv A B) (canIdToEquiv A B) =
--- -- substInv (Id U A B -> equiv A B)
--- -- (isEquiv (Id U A B) (equiv A B))
--- -- (canIdToEquiv A B) (transEquiv A B)
--- -- (canIdToEquivIsTransEquiv A B)
+-- -- univalenceJ (A B : U) : isEquiv (Path U A B) (equiv A B) (canPathToEquiv A B) =
+-- -- substInv (Path U A B -> equiv A B)
+-- -- (isEquiv (Path U A B) (equiv A B))
+-- -- (canPathToEquiv A B) (transEquiv A B)
+-- -- (canPathToEquivIsTransEquiv A B)
-- -- (transEquivIsEquiv A B)
isContrProp (A : U) (h : isContr A) : prop A = p
where
- p (a b : A) : Id A a b = <i> comp (<_> A) h.1 [ (i = 0) -> h.2 a, (i = 1) -> h.2 b ]
+ p (a b : A) : Path A a b = <i> comp (<_> A) h.1 [ (i = 0) -> h.2 a, (i = 1) -> h.2 b ]
idEquiv (A : U) : equiv A A = (idfun A,idIsEquiv A)
contrSinglEquiv (A B : U) (f : equiv A B) :
- Id ((X : U) * equiv X B) (B,idEquiv B) (A,f) = rem
+ Path ((X : U) * equiv X B) (B,idEquiv B) (A,f) = rem
where
rem1 : prop ((X : U) * equiv X B) = isContrProp ((X : U) * equiv X B) (lem1 B)
- rem : Id ((X : U) * equiv X B) (B,idEquiv B) (A,f) = rem1 (B,idEquiv B) (A,f)
+ rem : Path ((X : U) * equiv X B) (B,idEquiv B) (A,f) = rem1 (B,idEquiv B) (A,f)
elimEquiv (B : U) (P : (A : U) -> (A -> B) -> U) (d : P B (idfun B))
(A : U) (f : equiv A B) : P A f.1 = rem
where
T (z : (X : U) * equiv X B) : U = P z.1 z.2.1
- rem1 : Id ((X : U) * equiv X B) (B,idEquiv B) (A,f) = contrSinglEquiv A B f
+ rem1 : Path ((X : U) * equiv X B) (B,idEquiv B) (A,f) = contrSinglEquiv A B f
rem : P A f.1 = subst ((X : U) * equiv X B) T (B,idEquiv B) (A,f) rem1 d
elimIso (B : U) (Q : (A : U) -> (A -> B) -> (B -> A) -> U)