Skip to content

Commit 5a4482e

Browse files
jamesmckinnaandreasabel
authored andcommitted
[v2.1-rc1] fixes #2400: use explicit quantification instead (#2429)
* fixes #2400: use explicit quantification * fix knock-ons
1 parent 949e065 commit 5a4482e

File tree

2 files changed

+8
-8
lines changed

2 files changed

+8
-8
lines changed

src/Function/Endo/Setoid.agda

+4-4
Original file line numberDiff line numberDiff line change
@@ -58,7 +58,7 @@ private
5858
∘-isMagma : IsMagma _≈_ _∘_
5959
∘-isMagma = record
6060
{ isEquivalence = isEquivalence
61-
; ∙-cong = λ {_} {_} {_} {v} x≈y u≈v S.trans u≈v (cong v x≈y)
61+
; ∙-cong = λ {_} {_} {_} {k} f≈g h≈k x S.trans (h≈k _) (cong k (f≈g x))
6262
}
6363

6464
∘-magma : Magma (c ⊔ e) (c ⊔ e)
@@ -67,7 +67,7 @@ private
6767
∘-isSemigroup : IsSemigroup _≈_ _∘_
6868
∘-isSemigroup = record
6969
{ isMagma = ∘-isMagma
70-
; assoc = λ _ _ _ S.refl
70+
; assoc = λ _ _ _ _ S.refl
7171
}
7272

7373
∘-semigroup : Semigroup (c ⊔ e) (c ⊔ e)
@@ -76,7 +76,7 @@ private
7676
∘-id-isMonoid : IsMonoid _≈_ _∘_ id
7777
∘-id-isMonoid = record
7878
{ isSemigroup = ∘-isSemigroup
79-
; identity = (λ _ S.refl) , (λ _ S.refl)
79+
; identity = (λ _ _ S.refl) , (λ _ _ S.refl)
8080
}
8181

8282
∘-id-monoid : Monoid (c ⊔ e) (c ⊔ e)
@@ -112,6 +112,6 @@ module _ (f : Endo) where
112112
^-isMonoidHomomorphism : IsMonoidHomomorphism +-0-rawMonoid ∘-id-rawMonoid (f ^_)
113113
^-isMonoidHomomorphism = record
114114
{ isMagmaHomomorphism = ^-isMagmaHomomorphism
115-
; ε-homo = S.refl
115+
; ε-homo = λ _ S.refl
116116
}
117117

src/Function/Relation/Binary/Setoid/Equality.agda

+4-4
Original file line numberDiff line numberDiff line change
@@ -24,16 +24,16 @@ private
2424

2525
infix 4 _≈_
2626
_≈_ : (f g : Func From To) Set (a₁ ⊔ b₂)
27-
f ≈ g = {x : From.Carrier} f ⟨$⟩ x To.≈ g ⟨$⟩ x
27+
f ≈ g = x f ⟨$⟩ x To.≈ g ⟨$⟩ x
2828

2929
refl : Reflexive _≈_
30-
refl = To.refl
30+
refl _ = To.refl
3131

3232
sym : Symmetric _≈_
33-
sym = λ f≈g To.sym f≈g
33+
sym f≈g x = To.sym (f≈g x)
3434

3535
trans : Transitive _≈_
36-
trans = λ f≈g g≈h To.trans f≈g g≈h
36+
trans f≈g g≈h x = To.trans (f≈g x) (g≈h x)
3737

3838
isEquivalence : IsEquivalence _≈_
3939
isEquivalence = record -- need to η-expand else Agda gets confused

0 commit comments

Comments
 (0)