Skip to content

Commit

Permalink
Add neutral product laws
Browse files Browse the repository at this point in the history
  • Loading branch information
nickdrozd committed May 6, 2020
1 parent 695c6ce commit 942a2c7
Showing 1 changed file with 14 additions and 0 deletions.
14 changes: 14 additions & 0 deletions libs/contrib/Control/Algebra/Laws.idr
Original file line number Diff line number Diff line change
Expand Up @@ -143,6 +143,20 @@ cancelRight x y z p =
rewrite groupInverseIsInverseL x in
monoidNeutralIsNeutralL z

||| ab = 0 -> a = b^-1
neutralProductInverseR : VerifiedGroup ty => (a, b : ty) ->
a <+> b = A.neutral -> a = inverse b
neutralProductInverseR a b prf =
cancelRight b a (inverse b) $
trans prf $ sym $ groupInverseIsInverseR b

||| ab = 0 -> a^-1 = b
neutralProductInverseL : VerifiedGroup ty => (a, b : ty) ->
a <+> b = A.neutral -> inverse a = b
neutralProductInverseL a b prf =
cancelLeft a (inverse a) b $
trans (groupInverseIsInverseL a) $ sym prf

||| For any a and b, ax = b and ya = b have solutions.
latinSquareProperty : VerifiedGroup t => (a, b : t) ->
((x : t ** a <+> x = b),
Expand Down

0 comments on commit 942a2c7

Please sign in to comment.