Skip to content

Commit 0fd6d50

Browse files
committed
Add neutral product laws
1 parent feedb9a commit 0fd6d50

File tree

1 file changed

+14
-0
lines changed

1 file changed

+14
-0
lines changed

libs/contrib/Control/Algebra/Laws.idr

Lines changed: 14 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -143,6 +143,20 @@ cancelRight x y z p =
143143
rewrite groupInverseIsInverseL x in
144144
monoidNeutralIsNeutralL z
145145

146+
||| ab = 0 -> a = b^-1
147+
neutralProductInverseR : VerifiedGroup ty => (a, b : ty) ->
148+
a <+> b = A.neutral -> a = inverse b
149+
neutralProductInverseR a b prf =
150+
cancelRight b a (inverse b) $
151+
trans prf $ sym $ groupInverseIsInverseR b
152+
153+
||| ab = 0 -> a^-1 = b
154+
neutralProductInverseL : VerifiedGroup ty => (a, b : ty) ->
155+
a <+> b = A.neutral -> inverse a = b
156+
neutralProductInverseL a b prf =
157+
cancelLeft a (inverse a) b $
158+
trans (groupInverseIsInverseL a) $ sym prf
159+
146160
||| For any a and b, ax = b and ya = b have solutions.
147161
latinSquareProperty : VerifiedGroup t => (a, b : t) ->
148162
((x : t ** a <+> x = b),

0 commit comments

Comments
 (0)