Skip to content
This repository was archived by the owner on Nov 12, 2025. It is now read-only.

Commit 553abe6

Browse files
committed
Add GroupHomomorphism
1 parent 63bc631 commit 553abe6

File tree

2 files changed

+29
-0
lines changed

2 files changed

+29
-0
lines changed
Lines changed: 28 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,28 @@
1+
import Prelude.Algebra as A
2+
import Control.Algebra
3+
import Control.Algebra.Laws
4+
5+
||| A homomorphism is a mapping that preserves group structure.
6+
interface (VerifiedGroup a, VerifiedGroup b) => GroupHomomorphism a b where
7+
to : a -> b
8+
toGroup : (x, y : a) ->
9+
to (x <+> y) = (to x) <+> (to y)
10+
11+
||| Homomorphism preserves neutral.
12+
homoNeutral : GroupHomomorphism a b =>
13+
to (the a A.neutral) = the b A.neutral
14+
homoNeutral {a} =
15+
let ae = the a A.neutral in
16+
selfSquareId (to ae) $
17+
trans (sym $ toGroup ae ae) $
18+
cong {f = to} $ monoidNeutralIsNeutralL ae
19+
20+
||| Homomorphism preserves inverses.
21+
homoInverse : GroupHomomorphism a b => (x : a) ->
22+
the b (to $ inverse x) = the b (inverse $ to x)
23+
homoInverse {a} {b} x =
24+
cancelRight (to x) (to $ inverse x) (inverse $ to x) $
25+
trans (sym $ toGroup (inverse x) x) $
26+
trans (cong {f = to} $ groupInverseIsInverseR x) $
27+
trans (homoNeutral {a=a} {b=b}) $
28+
sym $ groupInverseIsInverseR (to x)

libs/contrib/contrib.ipkg

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -12,6 +12,7 @@ modules = CFFI
1212
, Control.Algebra.NumericImplementations
1313
, Control.Algebra.ComplexImplementations
1414
, Control.Algebra.VectorSpace
15+
, Control.Algebra.GroupHomomorphism
1516

1617
, Control.Delayed
1718

0 commit comments

Comments
 (0)