File tree 1 file changed +10
-9
lines changed
src/Function/Relation/Binary/Setoid
1 file changed +10
-9
lines changed Original file line number Diff line number Diff line change @@ -15,6 +15,8 @@ module Function.Relation.Binary.Setoid.Equality {a₁ a₂ b₁ b₂ : Level}
15
15
open import Function.Bundles using (Func; _⟨$⟩_)
16
16
open import Relation.Binary.Definitions
17
17
using (Reflexive; Symmetric; Transitive)
18
+ open import Relation.Binary.Structures
19
+ using (IsEquivalence)
18
20
19
21
private
20
22
module To = Setoid To
@@ -33,17 +35,16 @@ sym = λ f≈g → To.sym f≈g
33
35
trans : Transitive _≈_
34
36
trans = λ f≈g g≈h → To.trans f≈g g≈h
35
37
36
- setoid : Setoid _ _
37
- setoid = record
38
- { Carrier = Func From To
39
- ; _≈_ = _≈_
40
- ; isEquivalence = record -- need to η-expand else Agda gets confused
41
- { refl = λ {f} → refl {f}
42
- ; sym = λ {f} {g} → sym {f} {g}
43
- ; trans = λ {f} {g} {h} → trans {f} {g} {h}
44
- }
38
+ isEquivalence : IsEquivalence _≈_
39
+ isEquivalence = record -- need to η-expand else Agda gets confused
40
+ { refl = λ {f} → refl {f}
41
+ ; sym = λ {f} {g} → sym {f} {g}
42
+ ; trans = λ {f} {g} {h} → trans {f} {g} {h}
45
43
}
46
44
45
+ setoid : Setoid _ _
46
+ setoid = record { isEquivalence = isEquivalence }
47
+
47
48
-- most of the time, this infix version is nicer to use
48
49
infixr 9 _⇨_
49
50
_⇨_ : Setoid _ _
You can’t perform that action at this time.
0 commit comments