Skip to content

Commit 9b7d5f2

Browse files
committed
Fix compilation with agda/agda#7581
1 parent 5449cc2 commit 9b7d5f2

File tree

8 files changed

+16
-17
lines changed

8 files changed

+16
-17
lines changed

Cubical/Categories/Constructions/Slice/Base.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -46,7 +46,7 @@ open SliceHom public
4646
-- can probably replace these by showing that SliceOb is isomorphic to Sigma and
4747
-- that paths are isomorphic to Sigma? But sounds like that would need a lot of transp
4848

49-
SliceOb-≡-intro : {a b} {f g}
49+
SliceOb-≡-intro : {a b} {f : C [ a , c ]} {g : C [ b , c ]}
5050
(p : a ≡ b)
5151
PathP (λ i C [ p i , c ]) f g
5252
sliceob {a} f ≡ sliceob {b} g

Cubical/Categories/DistLatticeSheaf/ComparisonLemma.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -399,7 +399,7 @@ module _ (L : DistLattice ℓ) (C : Category ℓ' ℓ'') (limitC : Limits {ℓ}
399399
(NatTransCone _ _ _ F (idTrans _) x)
400400
⋆⟨ C ⟩ limOfArrows (FLimCone (α ∘ˡ i) _) (GLimCone (α ∘ˡ i) _)
401401
(↓nt (α ∘ˡ i) x)
402-
goal x = sym (limArrowUnique _ _ _ _ (isConeMorComp x))
402+
goal x = sym (limArrowUnique {C = C} _ _ _ _ (isConeMorComp x))
403403
∙ limArrowCompLimOfArrows _ _ _ _ _
404404

405405
nIso (η winv) (F , isSheafF) = isIsoΣPropCat _ _ _

Cubical/Categories/Instances/CommAlgebras.agda

Lines changed: 2 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -675,10 +675,9 @@ module PreSheafFromUniversalProp (C : Category ℓ ℓ') (P : ob C → Type ℓ)
675675
assocDiagPath : Forgetful ∘F (universalPShf ∘F D) ≡ 𝓖 ∘F D
676676
assocDiagPath = F-assoc
677677

678-
conePathPCR : PathP (λ i Cone (assocDiagPath i) (F-ob (Forgetful ∘F universalPShf) c))
678+
conePathPCR : PathP (λ i Cone (assocDiagPath i) (F-ob 𝓖 c))
679679
(F-cone Forgetful (F-cone universalPShf cc)) (F-cone 𝓖 cc)
680-
conePathPCR = conePathPDiag -- why does everything have to be explicit?
681-
(λ v _ (Forgetful ∘F universalPShf) .F-hom {x = c} {y = D .F-ob v} (cc .coneOut v))
680+
conePathPCR = conePathPDiag (λ v _ 𝓖 .F-hom (cc .coneOut v))
682681

683682

684683
toLimCone : isLimCone _ _ (F-cone 𝓖 cc)

Cubical/Categories/Limits/RightKan.agda

Lines changed: 7 additions & 7 deletions
Original file line numberDiff line numberDiff line change
@@ -106,9 +106,9 @@ module _ {ℓC ℓC' ℓM ℓM' ℓA ℓA' : Level}
106106
path v = (F-hom Ran f) ⋆⟨ A ⟩ (F-hom Ran g) ⋆⟨ A ⟩ (limOut (limitA (z ↓Diag) (T* z)) v)
107107
≡⟨ ⋆Assoc A _ _ _ ⟩
108108
(F-hom Ran f) ⋆⟨ A ⟩ ((F-hom Ran g) ⋆⟨ A ⟩ (limOut (limitA (z ↓Diag) (T* z)) v))
109-
≡⟨ cong (seq' A (F-hom Ran f)) (limArrowCommutes _ _ _ _) ⟩
109+
≡⟨ cong (seq' A (F-hom Ran f)) (limArrowCommutes (limitA _ _) _ _ _) ⟩
110110
(F-hom Ran f) ⋆⟨ A ⟩ limOut (limitA (y ↓Diag) (T* y)) (j g .F-ob v)
111-
≡⟨ limArrowCommutes _ _ _ _ ⟩
111+
≡⟨ limArrowCommutes (limitA _ _) _ _ _ ⟩
112112
limOut (limitA (x ↓Diag) (T* x)) (j f .F-ob (j g .F-ob v))
113113
≡⟨ RanConeTrans f g v ⟩
114114
coneOut (RanCone (f ⋆⟨ C ⟩ g)) v ∎
@@ -121,7 +121,7 @@ module _ {ℓC ℓC' ℓM ℓM' ℓA ℓA' : Level}
121121
Ran .F-hom (K .F-hom f) ⋆⟨ A ⟩ coneOut (RanCone (id C)) (v , id C)
122122
≡⟨ cong (λ g Ran .F-hom (K .F-hom f) ⋆⟨ A ⟩ g) (sym (RanConeRefl (v , id C))) ⟩
123123
Ran .F-hom (K .F-hom f) ⋆⟨ A ⟩ limOut (limitA ((K .F-ob v) ↓Diag) (T* (K .F-ob v))) (v , id C)
124-
≡⟨ limArrowCommutes _ _ _ _ ⟩
124+
≡⟨ limArrowCommutes (limitA _ _) _ _ _ ⟩
125125
coneOut (RanCone (K .F-hom f)) (v , id C)
126126
≡⟨ cong (λ g limOut (limitA ((K .F-ob u) ↓Diag) (T* (K .F-ob u))) (v , g))
127127
(⋆IdR C (K .F-hom f) ∙ sym (⋆IdL C (K .F-hom f))) ⟩
@@ -169,15 +169,15 @@ module _ {ℓC ℓC' ℓM ℓM' ℓA ℓA' : Level}
169169
(S .F-hom f ⋆⟨ A ⟩ N-ob σ y) ⋆⟨ A ⟩ limOut (limitA (y ↓Diag) (T* y)) (u , g)
170170
≡⟨ ⋆Assoc A _ _ _ ⟩
171171
S .F-hom f ⋆⟨ A ⟩ (N-ob σ y ⋆⟨ A ⟩ limOut (limitA (y ↓Diag) (T* y)) (u , g))
172-
≡⟨ cong (seq' A (S .F-hom f)) (limArrowCommutes _ _ _ _) ⟩
172+
≡⟨ cong (seq' A (S .F-hom f)) (limArrowCommutes (limitA _ _) _ _ _) ⟩
173173
S .F-hom f ⋆⟨ A ⟩ (S .F-hom g ⋆⟨ A ⟩ α .N-ob u)
174174
≡⟨ sym (⋆Assoc A _ _ _) ⟩
175175
(S .F-hom f ⋆⟨ A ⟩ S .F-hom g) ⋆⟨ A ⟩ α .N-ob u
176176
≡⟨ cong (λ h h ⋆⟨ A ⟩ α .N-ob u) (sym (S .F-seq _ _)) ⟩
177177
S .F-hom (f ⋆⟨ C ⟩ g) ⋆⟨ A ⟩ α .N-ob u
178-
≡⟨ sym (limArrowCommutes _ _ _ _) ⟩
178+
≡⟨ sym (limArrowCommutes (limitA _ _) _ _ _) ⟩
179179
N-ob σ x ⋆⟨ A ⟩ limOut (limitA (x ↓Diag) (T* x)) (u , f ⋆⟨ C ⟩ g)
180-
≡⟨ cong (seq' A (N-ob σ x)) (sym (limArrowCommutes _ _ _ _)) ⟩
180+
≡⟨ cong (seq' A (N-ob σ x)) (sym (limArrowCommutes (limitA _ _) _ _ _)) ⟩
181181
N-ob σ x ⋆⟨ A ⟩ (Ran .F-hom f ⋆⟨ A ⟩ limOut (limitA (y ↓Diag) (T* y)) (u , g))
182182
≡⟨ sym (⋆Assoc A _ _ _) ⟩
183183
(N-ob σ x ⋆⟨ A ⟩ Ran .F-hom f) ⋆⟨ A ⟩ limOut (limitA (y ↓Diag) (T* y)) (u , g) ∎
@@ -200,7 +200,7 @@ module _ {ℓC ℓC' ℓM ℓM' ℓA ℓA' : Level}
200200
S .F-hom (id C ⋆⟨ C ⟩ id C) ⋆⟨ A ⟩ α .N-ob u
201201
≡⟨ refl ⟩
202202
NatTransCone S α (F-ob K u) .coneOut (u , id C ⋆⟨ C ⟩ id C)
203-
≡⟨ sym (limArrowCommutes _ _ _ _) ⟩
203+
≡⟨ sym (limArrowCommutes (limitA _ _) _ _ _) ⟩
204204
limArrow (limitA ((K .F-ob u) ↓Diag) (T* (K .F-ob u))) _ (NatTransCone S α (F-ob K u))
205205
⋆⟨ A ⟩ limOut (limitA ((K .F-ob u) ↓Diag) (T* (K .F-ob u))) (u , id C ⋆⟨ C ⟩ id C) ∎
206206

Cubical/Categories/Presheaf/Properties.agda

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -329,7 +329,7 @@ module _ {ℓS : Level} (C : Category ℓ ℓ') (F : Functor (C ^op) (SET ℓS))
329329
refl≡comm : PathP (λ i (F ⟪ f ⟫) (eq i) ≡ (eq' i)) refl comm
330330
refl≡comm = isOfHLevel→isOfHLevelDep 1 (λ (v , w) snd (F ⟅ d ⟆) ((F ⟪ f ⟫) w) v) refl comm λ i (eq' i , eq i)
331331

332-
X'≡subst : PathP (λ i fst (P ⟅ c , eq i ⟆)) X' (subst _ eq X')
332+
X'≡subst : PathP (λ i fst (P ⟅ c , eq i ⟆)) X' (subst (λ v fst (P ⟅ c , v ⟆)) eq X')
333333
X'≡subst = transport-filler (λ i fst (P ⟅ c , eq i ⟆)) X'
334334

335335
-- bottom left of the commuting diagram
@@ -355,7 +355,7 @@ module _ {ℓS : Level} (C : Category ℓ ℓ') (F : Functor (C ^op) (SET ℓS))
355355
right i = (α ⟦ c , eq i ⟧) (X'≡subst i)
356356
where
357357
-- this is exactly the same as the one from before, can refactor?
358-
X'≡subst : PathP (λ i fst (P ⟅ c , eq i ⟆)) X' (subst _ eq X')
358+
X'≡subst : PathP (λ i fst (P ⟅ c , eq i ⟆)) X' (subst (λ v fst (P ⟅ c , v ⟆)) eq X')
359359
X'≡subst = transport-filler _ _
360360

361361
-- extracted out type since need to use in in 'left' body as well

Cubical/Categories/RezkCompletion/Construction.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -51,7 +51,7 @@ module RezkByYoneda (C : Category ℓ ℓ) where
5151
ToYonedaImage = ToEssentialImage _
5252

5353
isWeakEquivalenceToYonedaImage : isWeakEquivalence ToYonedaImage
54-
isWeakEquivalenceToYonedaImage .fullfaith = isFullyFaithfulToEssentialImage _ isFullyFaithfulYO
54+
isWeakEquivalenceToYonedaImage .fullfaith = isFullyFaithfulToEssentialImage YO isFullyFaithfulYO
5555
isWeakEquivalenceToYonedaImage .esssurj = isEssentiallySurjToEssentialImage YO
5656

5757
isRezkCompletionToYonedaImage : isRezkCompletion ToYonedaImage

Cubical/Tactics/CategorySolver/Solver.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -51,7 +51,7 @@ module Eval (𝓒 : Category ℓ ℓ') where
5151
solve : {A B} (e₁ e₂ : W𝓒 [ A , B ])
5252
eval e₁ ≡ eval e₂
5353
⟦ e₁ ⟧c ≡ ⟦ e₂ ⟧c
54-
solve {A}{B} e₁ e₂ p = cong ⟦_⟧c (isFaithfulPseudoYoneda _ _ _ _ lem) where
54+
solve {A}{B} e₁ e₂ p = cong ⟦_⟧c (isFaithfulPseudoYoneda {C = W𝓒} _ _ _ _ lem) where
5555
lem : 𝓘 ⟪ e₁ ⟫ ≡ 𝓘 ⟪ e₂ ⟫
5656
lem = transport
5757
(λ i Yo-YoSem-agree (~ i) ⟪ e₁ ⟫ ≡ Yo-YoSem-agree (~ i) ⟪ e₂ ⟫)

Cubical/Tactics/FunctorSolver/Solver.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -55,7 +55,7 @@ module Eval (𝓒 : Category ℓc ℓc') (𝓓 : Category ℓd ℓd') (𝓕 : F
5555
(p : Path _ (YoSem.semH ⟪ e ⟫) (YoSem.semH ⟪ e' ⟫))
5656
Path _ (TautoSem.semH ⟪ e ⟫) (TautoSem.semH ⟪ e' ⟫)
5757
solve {A}{B} e e' p =
58-
cong (TautoSem.semH .F-hom) (isFaithfulPseudoYoneda _ _ _ _ lem) where
58+
cong (TautoSem.semH .F-hom) (isFaithfulPseudoYoneda {C = Free𝓓} _ _ _ _ lem) where
5959
lem : Path _ (PsYo ⟪ e ⟫) (PsYo ⟪ e' ⟫)
6060
lem = transport
6161
(λ i Path _

0 commit comments

Comments
 (0)