Skip to content

Commit 8eca37c

Browse files
committed
Add Algebra laws
1 parent 7a378ab commit 8eca37c

File tree

2 files changed

+164
-0
lines changed

2 files changed

+164
-0
lines changed
Lines changed: 162 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,162 @@
1+
module Control.Algebra.Laws
2+
3+
import Prelude
4+
import Control.Algebra
5+
6+
%default total
7+
8+
-- Monoids
9+
10+
||| Inverses are unique.
11+
public export
12+
uniqueInverse : VMonoid ty => (x, y, z : ty) ->
13+
y <+> x = neutral {ty} -> x <+> z = neutral {ty} -> y = z
14+
uniqueInverse x y z p q =
15+
rewrite sym $ monoidNeutralIsNeutralL y in
16+
rewrite sym q in
17+
rewrite semigroupOpIsAssociative y x z in
18+
rewrite p in
19+
rewrite monoidNeutralIsNeutralR z in
20+
Refl
21+
22+
-- Groups
23+
24+
||| Only identity is self-squaring.
25+
public export
26+
selfSquareId : VGroup ty => (x : ty) ->
27+
x <+> x = x -> x = neutral {ty}
28+
29+
||| Inverse elements commute.
30+
public export
31+
inverseCommute : VGroup ty => (x, y : ty) ->
32+
y <+> x = neutral {ty} -> x <+> y = neutral {ty}
33+
34+
||| Every element has a right inverse.
35+
public export
36+
groupInverseIsInverseL : VGroup ty => (x : ty) ->
37+
x <+> inverse x = neutral {ty}
38+
39+
||| -(-x) = x
40+
public export
41+
inverseSquaredIsIdentity : VGroup ty => (x : ty) ->
42+
inverse (inverse x) = x
43+
44+
||| If every square in a group is identity, the group is commutative.
45+
public export
46+
squareIdCommutative : VGroup ty => (x, y : ty) ->
47+
((a : ty) -> a <+> a = neutral {ty}) ->
48+
x <+> y = y <+> x
49+
50+
||| -0 = 0
51+
public export
52+
inverseNeutralIsNeutral : VGroup ty =>
53+
inverse (neutral {ty}) = neutral {ty}
54+
55+
||| -(x + y) = -y + -x
56+
public export
57+
inverseOfSum : VGroup ty => (l, r : ty) ->
58+
inverse (l <+> r) = inverse r <+> inverse l
59+
60+
||| y = z if x + y = x + z
61+
public export
62+
cancelLeft : VGroup ty => (x, y, z : ty) ->
63+
x <+> y = x <+> z -> y = z
64+
cancelLeft x y z p =
65+
rewrite sym $ monoidNeutralIsNeutralR y in
66+
rewrite sym $ groupInverseIsInverseR x in
67+
rewrite sym $ semigroupOpIsAssociative (inverse x) x y in
68+
rewrite p in
69+
rewrite semigroupOpIsAssociative (inverse x) x z in
70+
rewrite groupInverseIsInverseR x in
71+
monoidNeutralIsNeutralR z
72+
73+
||| y = z if y + x = z + x.
74+
public export
75+
cancelRight : VGroup ty => (x, y, z : ty) ->
76+
y <+> x = z <+> x -> y = z
77+
78+
||| ab = 0 -> a = b^-1
79+
public export
80+
neutralProductInverseR : VGroup ty => (a, b : ty) ->
81+
a <+> b = neutral {ty} -> a = inverse b
82+
83+
||| ab = 0 -> a^-1 = b
84+
public export
85+
neutralProductInverseL : VGroup ty => (a, b : ty) ->
86+
a <+> b = neutral {ty} -> inverse a = b
87+
neutralProductInverseL a b prf =
88+
cancelLeft a (inverse a) b $
89+
trans (groupInverseIsInverseL a) $ sym prf
90+
91+
||| For any a and b, ax = b and ya = b have solutions.
92+
public export
93+
latinSquareProperty : VGroup ty => (a, b : ty) ->
94+
((x : ty ** a <+> x = b),
95+
(y : ty ** y <+> a = b))
96+
97+
||| For any a, b, x, the solution to ax = b is unique.
98+
public export
99+
uniqueSolutionR : VGroup ty => (a, b, x, y : ty) ->
100+
a <+> x = b -> a <+> y = b -> x = y
101+
uniqueSolutionR a b x y p q = cancelLeft a x y $ trans p (sym q)
102+
103+
||| For any a, b, y, the solution to ya = b is unique.
104+
public export
105+
uniqueSolutionL : VGroup t => (a, b, x, y : t) ->
106+
x <+> a = b -> y <+> a = b -> x = y
107+
uniqueSolutionL a b x y p q = cancelRight a x y $ trans p (sym q)
108+
109+
||| -(x + y) = -x + -y
110+
public export
111+
inverseDistributesOverGroupOp : AbelianGroup ty => (l, r : ty) ->
112+
inverse (l <+> r) = inverse l <+> inverse r
113+
inverseDistributesOverGroupOp l r =
114+
rewrite groupOpIsCommutative (inverse l) (inverse r) in
115+
inverseOfSum l r
116+
117+
||| Homomorphism preserves neutral.
118+
public export
119+
homoNeutral : GroupHomomorphism a b =>
120+
to (neutral {ty=a}) = neutral {ty=b}
121+
122+
||| Homomorphism preserves inverses.
123+
public export
124+
homoInverse : GroupHomomorphism a b => (x : a) ->
125+
the b (to $ inverse x) = the b (inverse $ to x)
126+
127+
-- Rings
128+
129+
||| 0x = x
130+
public export
131+
multNeutralAbsorbingL : VRing ty => (r : ty) ->
132+
neutral {ty} <.> r = neutral {ty}
133+
134+
||| x0 = 0
135+
public export
136+
multNeutralAbsorbingR : VRing ty => (l : ty) ->
137+
l <.> neutral {ty} = neutral {ty}
138+
139+
||| (-x)y = -(xy)
140+
public export
141+
multInverseInversesL : VRing ty => (l, r : ty) ->
142+
inverse l <.> r = inverse (l <.> r)
143+
144+
||| x(-y) = -(xy)
145+
public export
146+
multInverseInversesR : VRing ty => (l, r : ty) ->
147+
l <.> inverse r = inverse (l <.> r)
148+
149+
||| (-x)(-y) = xy
150+
public export
151+
multNegativeByNegativeIsPositive : VRing ty => (l, r : ty) ->
152+
inverse l <.> inverse r = l <.> r
153+
154+
||| (-1)x = -x
155+
public export
156+
inverseOfUnityR : VRingWithUnity ty => (x : ty) ->
157+
inverse (unity {ty}) <.> x = inverse x
158+
159+
||| x(-1) = -x
160+
public export
161+
inverseOfUnityL : VRingWithUnity ty => (x : ty) ->
162+
x <.> inverse (unity {ty}) = inverse x

libs/contrib/contrib.ipkg

Lines changed: 2 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1,7 +1,9 @@
11
package contrib
22

33
modules = Control.Delayed,
4+
45
Control.Algebra,
6+
Control.Algebra.Laws,
57

68
Data.List.TailRec,
79
Data.List.Equalities,

0 commit comments

Comments
 (0)