Skip to content

Commit 9d1bb82

Browse files
authored
Merge pull request #242 from nickdrozd/algebra
Add Algebra interfaces and laws
2 parents 2321513 + 028a1cb commit 9d1bb82

File tree

3 files changed

+448
-0
lines changed

3 files changed

+448
-0
lines changed

libs/contrib/Control/Algebra.idr

Lines changed: 154 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,154 @@
1+
module Control.Algebra
2+
3+
infixl 6 <->
4+
infixl 7 <.>
5+
6+
public export
7+
interface Semigroup ty => SemigroupV ty where
8+
semigroupOpIsAssociative : (l, c, r : ty) ->
9+
l <+> (c <+> r) = (l <+> c) <+> r
10+
11+
public export
12+
interface (Monoid ty, SemigroupV ty) => MonoidV ty where
13+
monoidNeutralIsNeutralL : (l : ty) -> l <+> neutral {ty} = l
14+
monoidNeutralIsNeutralR : (r : ty) -> neutral {ty} <+> r = r
15+
16+
||| Sets equipped with a single binary operation that is associative,
17+
||| along with a neutral element for that binary operation and
18+
||| inverses for all elements. Satisfies the following laws:
19+
|||
20+
||| + Associativity of `<+>`:
21+
||| forall a b c, a <+> (b <+> c) == (a <+> b) <+> c
22+
||| + Neutral for `<+>`:
23+
||| forall a, a <+> neutral == a
24+
||| forall a, neutral <+> a == a
25+
||| + Inverse for `<+>`:
26+
||| forall a, a <+> inverse a == neutral
27+
||| forall a, inverse a <+> a == neutral
28+
public export
29+
interface MonoidV ty => Group ty where
30+
inverse : ty -> ty
31+
32+
groupInverseIsInverseR : (r : ty) -> inverse r <+> r = neutral {ty}
33+
34+
(<->) : Group ty => ty -> ty -> ty
35+
(<->) left right = left <+> (inverse right)
36+
37+
||| Sets equipped with a single binary operation that is associative
38+
||| and commutative, along with a neutral element for that binary
39+
||| operation and inverses for all elements. Satisfies the following
40+
||| laws:
41+
|||
42+
||| + Associativity of `<+>`:
43+
||| forall a b c, a <+> (b <+> c) == (a <+> b) <+> c
44+
||| + Commutativity of `<+>`:
45+
||| forall a b, a <+> b == b <+> a
46+
||| + Neutral for `<+>`:
47+
||| forall a, a <+> neutral == a
48+
||| forall a, neutral <+> a == a
49+
||| + Inverse for `<+>`:
50+
||| forall a, a <+> inverse a == neutral
51+
||| forall a, inverse a <+> a == neutral
52+
public export
53+
interface Group ty => AbelianGroup ty where
54+
groupOpIsCommutative : (l, r : ty) -> l <+> r = r <+> l
55+
56+
||| A homomorphism is a mapping that preserves group structure.
57+
public export
58+
interface (Group a, Group b) => GroupHomomorphism a b where
59+
to : a -> b
60+
61+
toGroup : (x, y : a) ->
62+
to (x <+> y) = (to x) <+> (to y)
63+
64+
||| Sets equipped with two binary operations, one associative and
65+
||| commutative supplied with a neutral element, and the other
66+
||| associative, with distributivity laws relating the two operations.
67+
||| Satisfies the following laws:
68+
|||
69+
||| + Associativity of `<+>`:
70+
||| forall a b c, a <+> (b <+> c) == (a <+> b) <+> c
71+
||| + Commutativity of `<+>`:
72+
||| forall a b, a <+> b == b <+> a
73+
||| + Neutral for `<+>`:
74+
||| forall a, a <+> neutral == a
75+
||| forall a, neutral <+> a == a
76+
||| + Inverse for `<+>`:
77+
||| forall a, a <+> inverse a == neutral
78+
||| forall a, inverse a <+> a == neutral
79+
||| + Associativity of `<.>`:
80+
||| forall a b c, a <.> (b <.> c) == (a <.> b) <.> c
81+
||| + Distributivity of `<.>` and `<+>`:
82+
||| forall a b c, a <.> (b <+> c) == (a <.> b) <+> (a <.> c)
83+
||| forall a b c, (a <+> b) <.> c == (a <.> c) <+> (b <.> c)
84+
public export
85+
interface Group ty => Ring ty where
86+
(<.>) : ty -> ty -> ty
87+
88+
ringOpIsAssociative : (l, c, r : ty) ->
89+
l <.> (c <.> r) = (l <.> c) <.> r
90+
ringOpIsDistributiveL : (l, c, r : ty) ->
91+
l <.> (c <+> r) = (l <.> c) <+> (l <.> r)
92+
ringOpIsDistributiveR : (l, c, r : ty) ->
93+
(l <+> c) <.> r = (l <.> r) <+> (c <.> r)
94+
95+
||| Sets equipped with two binary operations, one associative and
96+
||| commutative supplied with a neutral element, and the other
97+
||| associative supplied with a neutral element, with distributivity
98+
||| laws relating the two operations. Satisfies the following laws:
99+
|||
100+
||| + Associativity of `<+>`:
101+
||| forall a b c, a <+> (b <+> c) == (a <+> b) <+> c
102+
||| + Commutativity of `<+>`:
103+
||| forall a b, a <+> b == b <+> a
104+
||| + Neutral for `<+>`:
105+
||| forall a, a <+> neutral == a
106+
||| forall a, neutral <+> a == a
107+
||| + Inverse for `<+>`:
108+
||| forall a, a <+> inverse a == neutral
109+
||| forall a, inverse a <+> a == neutral
110+
||| + Associativity of `<.>`:
111+
||| forall a b c, a <.> (b <.> c) == (a <.> b) <.> c
112+
||| + Neutral for `<.>`:
113+
||| forall a, a <.> unity == a
114+
||| forall a, unity <.> a == a
115+
||| + Distributivity of `<.>` and `<+>`:
116+
||| forall a b c, a <.> (b <+> c) == (a <.> b) <+> (a <.> c)
117+
||| forall a b c, (a <+> b) <.> c == (a <.> c) <+> (b <.> c)
118+
public export
119+
interface Ring ty => RingWithUnity ty where
120+
unity : ty
121+
122+
unityIsRingIdL : (l : ty) -> l <.> unity = l
123+
unityIsRingIdR : (r : ty) -> unity <.> r = r
124+
125+
||| Sets equipped with two binary operations – both associative,
126+
||| commutative and possessing a neutral element – and distributivity
127+
||| laws relating the two operations. All elements except the additive
128+
||| identity should have a multiplicative inverse. Should (but may
129+
||| not) satisfy the following laws:
130+
|||
131+
||| + Associativity of `<+>`:
132+
||| forall a b c, a <+> (b <+> c) == (a <+> b) <+> c
133+
||| + Commutativity of `<+>`:
134+
||| forall a b, a <+> b == b <+> a
135+
||| + Neutral for `<+>`:
136+
||| forall a, a <+> neutral == a
137+
||| forall a, neutral <+> a == a
138+
||| + Inverse for `<+>`:
139+
||| forall a, a <+> inverse a == neutral
140+
||| forall a, inverse a <+> a == neutral
141+
||| + Associativity of `<.>`:
142+
||| forall a b c, a <.> (b <.> c) == (a <.> b) <.> c
143+
||| + Unity for `<.>`:
144+
||| forall a, a <.> unity == a
145+
||| forall a, unity <.> a == a
146+
||| + InverseM of `<.>`, except for neutral
147+
||| forall a /= neutral, a <.> inverseM a == unity
148+
||| forall a /= neutral, inverseM a <.> a == unity
149+
||| + Distributivity of `<.>` and `<+>`:
150+
||| forall a b c, a <.> (b <+> c) == (a <.> b) <+> (a <.> c)
151+
||| forall a b c, (a <+> b) <.> c == (a <.> c) <+> (b <.> c)
152+
public export
153+
interface RingWithUnity ty => Field ty where
154+
inverseM : (x : ty) -> Not (x = neutral {ty}) -> ty

0 commit comments

Comments
 (0)