File tree Expand file tree Collapse file tree 2 files changed +64
-6
lines changed
src/Function/Relation/Binary/Setoid Expand file tree Collapse file tree 2 files changed +64
-6
lines changed Original file line number Diff line number Diff line change @@ -48,18 +48,26 @@ Deprecated names
4848
4949New modules
5050-----------
51-
52- * Symmetric interior of a binary relation
53- ```
54- Relation.Binary.Construct.Interior.Symmetric
55- ```
56-
5751* ` Algebra.Module.Bundles.Raw ` : raw bundles for module-like algebraic structures
5852
5953* Nagata's construction of the "idealization of a module":
6054 ``` agda
6155 Algebra.Module.Construct.Idealization
6256 ```
57+
58+ * ` Function.Relation.Binary.Equality `
59+ ``` agda
60+ setoid : Setoid a₁ a₂ → Setoid b₁ b₂ → Setoid _ _
61+ ```
62+ and a convenient infix version
63+ ``` agda
64+ _⇨_ = setoid
65+ ```
66+
67+ * Symmetric interior of a binary relation
68+ ```
69+ Relation.Binary.Construct.Interior.Symmetric
70+ ```
6371
6472Additions to existing modules
6573-----------------------------
Original file line number Diff line number Diff line change 1+ ------------------------------------------------------------------------
2+ -- The Agda standard library
3+ --
4+ -- Function Equality setoid
5+ ------------------------------------------------------------------------
6+
7+ {-# OPTIONS --cubical-compatible --safe #-}
8+
9+ open import Level using (Level; _⊔_)
10+ open import Relation.Binary.Bundles using (Setoid)
11+
12+ module Function.Relation.Binary.Setoid.Equality {a₁ a₂ b₁ b₂ : Level}
13+ (From : Setoid a₁ a₂) (To : Setoid b₁ b₂) where
14+
15+ open import Function.Bundles using (Func; _⟨$⟩_)
16+ open import Relation.Binary.Definitions
17+ using (Reflexive; Symmetric; Transitive)
18+
19+ private
20+ module To = Setoid To
21+ module From = Setoid From
22+
23+ infix 4 _≈_
24+ _≈_ : (f g : Func From To) → Set (a₁ ⊔ b₂)
25+ f ≈ g = {x : From.Carrier} → f ⟨$⟩ x To.≈ g ⟨$⟩ x
26+
27+ refl : Reflexive _≈_
28+ refl = To.refl
29+
30+ sym : Symmetric _≈_
31+ sym = λ f≈g → To.sym f≈g
32+
33+ trans : Transitive _≈_
34+ trans = λ f≈g g≈h → To.trans f≈g g≈h
35+
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+ }
45+ }
46+
47+ -- most of the time, this infix version is nicer to use
48+ infixr 9 _⇨_
49+ _⇨_ : Setoid _ _
50+ _⇨_ = setoid
You can’t perform that action at this time.
0 commit comments