Skip to content

Commit 2618431

Browse files
authored
[ add ] product structure on RawSetoid (#2720)
* [ add ] product structure on `RawSetoid` * fix: export parametrisation * fix: eta contract definitions * fix: full stops * fix: redundant comments/imports * unbundle definitions; leave exports bundled * fix: cosmetic comment to trigger fresh test run
1 parent ca27adf commit 2618431

File tree

2 files changed

+84
-1
lines changed

2 files changed

+84
-1
lines changed

CHANGELOG.md

Lines changed: 3 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -163,7 +163,9 @@ New modules
163163

164164
* `Data.List.Relation.Binary.Suffix.Propositional.Properties` showing the equivalence to right divisibility induced by the list monoid.
165165

166-
* `Data.Sign.Show` to show a sign
166+
* `Data.Sign.Show` to show a sign.
167+
168+
* `Relation.Binary.Morphism.Construct.Product` to plumb in the (categorical) product structure on `RawSetoid`.
167169

168170
* `Relation.Binary.Properties.PartialSetoid` to systematise properties of a PER
169171

Lines changed: 81 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,81 @@
1+
------------------------------------------------------------------------
2+
-- The Agda standard library
3+
--
4+
-- The projection morphisms for relational structures arising from the
5+
-- non-dependent product construction
6+
------------------------------------------------------------------------
7+
8+
{-# OPTIONS --safe --cubical-compatible #-}
9+
10+
module Relation.Binary.Morphism.Construct.Product where
11+
12+
import Data.Product.Base as Product using (<_,_>; proj₁; proj₂)
13+
open import Data.Product.Relation.Binary.Pointwise.NonDependent as Pointwise
14+
using (Pointwise)
15+
open import Level using (Level)
16+
open import Relation.Binary.Bundles.Raw using (RawSetoid)
17+
open import Relation.Binary.Core using (Rel)
18+
open import Relation.Binary.Morphism.Structures using (IsRelHomomorphism)
19+
20+
private
21+
variable
22+
a b c ℓ₁ ℓ₂ ℓ : Level
23+
A : Set a
24+
B : Set b
25+
C : Set c
26+
27+
28+
------------------------------------------------------------------------
29+
-- definitions
30+
31+
module _ (_≈₁_ : Rel A ℓ₁) (_≈₂_ : Rel B ℓ₂) where
32+
33+
private
34+
35+
_≈_ = Pointwise _≈₁_ _≈₂_
36+
37+
module Proj₁ where
38+
39+
isRelHomomorphism : IsRelHomomorphism _≈_ _≈₁_ Product.proj₁
40+
isRelHomomorphism = record { cong = Product.proj₁ }
41+
42+
43+
module Proj₂ where
44+
45+
isRelHomomorphism : IsRelHomomorphism _≈_ _≈₂_ Product.proj₂
46+
isRelHomomorphism = record { cong = Product.proj₂ }
47+
48+
49+
module Pair (_≈′_ : Rel C ℓ) where
50+
51+
isRelHomomorphism : {h k}
52+
IsRelHomomorphism _≈′_ _≈₁_ h
53+
IsRelHomomorphism _≈′_ _≈₂_ k
54+
IsRelHomomorphism _≈′_ _≈_ Product.< h , k >
55+
isRelHomomorphism H K = record { cong = Product.< H.cong , K.cong > }
56+
where
57+
module H = IsRelHomomorphism H
58+
module K = IsRelHomomorphism K
59+
60+
61+
------------------------------------------------------------------------
62+
-- package up for export
63+
64+
module _ {M : RawSetoid a ℓ₁} {N : RawSetoid b ℓ₂} where
65+
66+
private
67+
68+
module M = RawSetoid M
69+
module N = RawSetoid N
70+
71+
proj₁ = Proj₁.isRelHomomorphism M._≈_ N._≈_
72+
proj₂ = Proj₂.isRelHomomorphism M._≈_ N._≈_
73+
74+
module _ {P : RawSetoid c ℓ} where
75+
76+
private
77+
78+
module P = RawSetoid P
79+
80+
<_,_> = Pair.isRelHomomorphism M._≈_ N._≈_ P._≈_
81+

0 commit comments

Comments
 (0)