Skip to content

Commit ebe659c

Browse files
authored
Adds Algebra.Morphism.Construct.DirectProduct (#2715)
* Adds Algebra.Morphism.Construct.DirectProduct * Update src/Algebra/Morphism/Construct/DirectProduct.agda * Address comments * Add module to reexport with implicit arguments * Removed unused private module * Rename H to G * Rename h to g
1 parent a53588f commit ebe659c

File tree

2 files changed

+144
-0
lines changed

2 files changed

+144
-0
lines changed

CHANGELOG.md

Lines changed: 2 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -161,6 +161,8 @@ New modules
161161

162162
* `Algebra.Module.Properties.{Bimodule|LeftModule|RightModule}`.
163163

164+
* `Algebra.Morphism.Construct.DirectProduct`.
165+
164166
* `Data.List.Base.{and|or|any|all}` have been lifted out into `Data.Bool.ListAction`.
165167

166168
* `Data.List.Base.{sum|product}` and their properties have been lifted out into `Data.Nat.ListAction` and `Data.Nat.ListAction.Properties`.
Lines changed: 142 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,142 @@
1+
------------------------------------------------------------------------
2+
-- The Agda standard library
3+
--
4+
-- The projection morphisms for algebraic structures arising from the
5+
-- direct product construction
6+
------------------------------------------------------------------------
7+
8+
{-# OPTIONS --safe --cubical-compatible #-}
9+
10+
module Algebra.Morphism.Construct.DirectProduct where
11+
12+
open import Algebra.Bundles using (RawMagma; RawMonoid)
13+
open import Algebra.Construct.DirectProduct using (rawMagma; rawMonoid)
14+
open import Algebra.Morphism.Structures
15+
using ( module MagmaMorphisms
16+
; module MonoidMorphisms
17+
)
18+
open import Data.Product as Product
19+
using (_,_)
20+
open import Level using (Level)
21+
open import Relation.Binary.Definitions using (Reflexive)
22+
import Relation.Binary.Morphism.Construct.Product as RP
23+
24+
private
25+
variable
26+
a b c ℓ₁ ℓ₂ ℓ₃ : Level
27+
28+
------------------------------------------------------------------------
29+
-- Magmas
30+
31+
module Magma (M : RawMagma a ℓ₁) (N : RawMagma b ℓ₂) where
32+
open MagmaMorphisms
33+
34+
private
35+
module M = RawMagma M
36+
module N = RawMagma N
37+
38+
module Proj₁ (refl : Reflexive M._≈_) where
39+
40+
isMagmaHomomorphism : IsMagmaHomomorphism (rawMagma M N) M Product.proj₁
41+
isMagmaHomomorphism = record
42+
{ isRelHomomorphism = RP.proj₁
43+
; homo = λ _ _ refl
44+
}
45+
46+
module Proj₂ (refl : Reflexive N._≈_) where
47+
48+
isMagmaHomomorphism : IsMagmaHomomorphism (rawMagma M N) N Product.proj₂
49+
isMagmaHomomorphism = record
50+
{ isRelHomomorphism = RP.proj₂
51+
; homo = λ _ _ refl
52+
}
53+
54+
module Pair (P : RawMagma c ℓ₃) where
55+
56+
isMagmaHomomorphism : {f g}
57+
IsMagmaHomomorphism P M f
58+
IsMagmaHomomorphism P N g
59+
IsMagmaHomomorphism P (rawMagma M N) (Product.< f , g >)
60+
isMagmaHomomorphism F G = record
61+
{ isRelHomomorphism = RP.< F.isRelHomomorphism , G.isRelHomomorphism >
62+
; homo = λ x y F.homo x y , G.homo x y
63+
}
64+
where
65+
module F = IsMagmaHomomorphism F
66+
module G = IsMagmaHomomorphism G
67+
68+
-- Package for export
69+
module Magma-Export {M : RawMagma a ℓ₁} {N : RawMagma b ℓ₂} where
70+
open Magma
71+
72+
private
73+
module M = RawMagma M
74+
module N = RawMagma N
75+
76+
module _ {refl : Reflexive M._≈_} where
77+
proj₁ = Proj₁.isMagmaHomomorphism M M refl
78+
79+
module _ {refl : Reflexive N._≈_} where
80+
proj₂ = Proj₂.isMagmaHomomorphism M N refl
81+
82+
module _ {P : RawMagma c ℓ₃} where
83+
<_,_> = Pair.isMagmaHomomorphism M N P
84+
85+
------------------------------------------------------------------------
86+
-- Monoids
87+
88+
module Monoid (M : RawMonoid a ℓ₁) (N : RawMonoid b ℓ₂) where
89+
open MonoidMorphisms
90+
91+
private
92+
module M = RawMonoid M
93+
module N = RawMonoid N
94+
95+
module Proj₁ (refl : Reflexive M._≈_) where
96+
97+
isMonoidHomomorphism : IsMonoidHomomorphism (rawMonoid M N) M Product.proj₁
98+
isMonoidHomomorphism = record
99+
{ isMagmaHomomorphism = Magma.Proj₁.isMagmaHomomorphism M.rawMagma N.rawMagma refl
100+
; ε-homo = refl
101+
}
102+
103+
module Proj₂ (refl : Reflexive N._≈_) where
104+
105+
isMonoidHomomorphism : IsMonoidHomomorphism (rawMonoid M N) N Product.proj₂
106+
isMonoidHomomorphism = record
107+
{ isMagmaHomomorphism = Magma.Proj₂.isMagmaHomomorphism M.rawMagma N.rawMagma refl
108+
; ε-homo = refl
109+
}
110+
111+
module Pair (P : RawMonoid c ℓ₃) where
112+
113+
private
114+
module P = RawMonoid P
115+
116+
isMonoidHomomorphism : {f g}
117+
IsMonoidHomomorphism P M f
118+
IsMonoidHomomorphism P N g
119+
IsMonoidHomomorphism P (rawMonoid M N) (Product.< f , g >)
120+
isMonoidHomomorphism F G = record
121+
{ isMagmaHomomorphism = Magma.Pair.isMagmaHomomorphism M.rawMagma N.rawMagma P.rawMagma F.isMagmaHomomorphism G.isMagmaHomomorphism
122+
; ε-homo = F.ε-homo , G.ε-homo }
123+
where
124+
module F = IsMonoidHomomorphism F
125+
module G = IsMonoidHomomorphism G
126+
127+
-- Package for export
128+
module Monoid-Export {M : RawMonoid a ℓ₁} {N : RawMonoid b ℓ₂} where
129+
open Monoid
130+
131+
private
132+
module M = RawMonoid M
133+
module N = RawMonoid N
134+
135+
module _ {refl : Reflexive M._≈_} where
136+
proj₁ = Proj₁.isMonoidHomomorphism M M refl
137+
138+
module _ {refl : Reflexive N._≈_} where
139+
proj₂ = Proj₂.isMonoidHomomorphism M N refl
140+
141+
module _ {P : RawMonoid c ℓ₃} where
142+
<_,_> = Pair.isMonoidHomomorphism M N P

0 commit comments

Comments
 (0)