Skip to content

Commit

Permalink
Add Algebra laws
Browse files Browse the repository at this point in the history
  • Loading branch information
nickdrozd committed Jun 5, 2020
1 parent 7a378ab commit 8eca37c
Show file tree
Hide file tree
Showing 2 changed files with 164 additions and 0 deletions.
162 changes: 162 additions & 0 deletions libs/contrib/Control/Algebra/Laws.idr
Original file line number Diff line number Diff line change
@@ -0,0 +1,162 @@
module Control.Algebra.Laws

import Prelude
import Control.Algebra

%default total

-- Monoids

||| Inverses are unique.
public export
uniqueInverse : VMonoid ty => (x, y, z : ty) ->
y <+> x = neutral {ty} -> x <+> z = neutral {ty} -> y = z
uniqueInverse x y z p q =
rewrite sym $ monoidNeutralIsNeutralL y in
rewrite sym q in
rewrite semigroupOpIsAssociative y x z in
rewrite p in
rewrite monoidNeutralIsNeutralR z in
Refl

-- Groups

||| Only identity is self-squaring.
public export
selfSquareId : VGroup ty => (x : ty) ->
x <+> x = x -> x = neutral {ty}

||| Inverse elements commute.
public export
inverseCommute : VGroup ty => (x, y : ty) ->
y <+> x = neutral {ty} -> x <+> y = neutral {ty}

||| Every element has a right inverse.
public export
groupInverseIsInverseL : VGroup ty => (x : ty) ->
x <+> inverse x = neutral {ty}

||| -(-x) = x
public export
inverseSquaredIsIdentity : VGroup ty => (x : ty) ->
inverse (inverse x) = x

||| If every square in a group is identity, the group is commutative.
public export
squareIdCommutative : VGroup ty => (x, y : ty) ->
((a : ty) -> a <+> a = neutral {ty}) ->
x <+> y = y <+> x

||| -0 = 0
public export
inverseNeutralIsNeutral : VGroup ty =>
inverse (neutral {ty}) = neutral {ty}

||| -(x + y) = -y + -x
public export
inverseOfSum : VGroup ty => (l, r : ty) ->
inverse (l <+> r) = inverse r <+> inverse l

||| y = z if x + y = x + z
public export
cancelLeft : VGroup ty => (x, y, z : ty) ->
x <+> y = x <+> z -> y = z
cancelLeft x y z p =
rewrite sym $ monoidNeutralIsNeutralR y in
rewrite sym $ groupInverseIsInverseR x in
rewrite sym $ semigroupOpIsAssociative (inverse x) x y in
rewrite p in
rewrite semigroupOpIsAssociative (inverse x) x z in
rewrite groupInverseIsInverseR x in
monoidNeutralIsNeutralR z

||| y = z if y + x = z + x.
public export
cancelRight : VGroup ty => (x, y, z : ty) ->
y <+> x = z <+> x -> y = z

||| ab = 0 -> a = b^-1
public export
neutralProductInverseR : VGroup ty => (a, b : ty) ->
a <+> b = neutral {ty} -> a = inverse b

||| ab = 0 -> a^-1 = b
public export
neutralProductInverseL : VGroup ty => (a, b : ty) ->
a <+> b = neutral {ty} -> 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.
public export
latinSquareProperty : VGroup ty => (a, b : ty) ->
((x : ty ** a <+> x = b),
(y : ty ** y <+> a = b))

||| For any a, b, x, the solution to ax = b is unique.
public export
uniqueSolutionR : VGroup ty => (a, b, x, y : ty) ->
a <+> x = b -> a <+> y = b -> x = y
uniqueSolutionR a b x y p q = cancelLeft a x y $ trans p (sym q)

||| For any a, b, y, the solution to ya = b is unique.
public export
uniqueSolutionL : VGroup t => (a, b, x, y : t) ->
x <+> a = b -> y <+> a = b -> x = y
uniqueSolutionL a b x y p q = cancelRight a x y $ trans p (sym q)

||| -(x + y) = -x + -y
public export
inverseDistributesOverGroupOp : AbelianGroup ty => (l, r : ty) ->
inverse (l <+> r) = inverse l <+> inverse r
inverseDistributesOverGroupOp l r =
rewrite groupOpIsCommutative (inverse l) (inverse r) in
inverseOfSum l r

||| Homomorphism preserves neutral.
public export
homoNeutral : GroupHomomorphism a b =>
to (neutral {ty=a}) = neutral {ty=b}

||| Homomorphism preserves inverses.
public export
homoInverse : GroupHomomorphism a b => (x : a) ->
the b (to $ inverse x) = the b (inverse $ to x)

-- Rings

||| 0x = x
public export
multNeutralAbsorbingL : VRing ty => (r : ty) ->
neutral {ty} <.> r = neutral {ty}

||| x0 = 0
public export
multNeutralAbsorbingR : VRing ty => (l : ty) ->
l <.> neutral {ty} = neutral {ty}

||| (-x)y = -(xy)
public export
multInverseInversesL : VRing ty => (l, r : ty) ->
inverse l <.> r = inverse (l <.> r)

||| x(-y) = -(xy)
public export
multInverseInversesR : VRing ty => (l, r : ty) ->
l <.> inverse r = inverse (l <.> r)

||| (-x)(-y) = xy
public export
multNegativeByNegativeIsPositive : VRing ty => (l, r : ty) ->
inverse l <.> inverse r = l <.> r

||| (-1)x = -x
public export
inverseOfUnityR : VRingWithUnity ty => (x : ty) ->
inverse (unity {ty}) <.> x = inverse x

||| x(-1) = -x
public export
inverseOfUnityL : VRingWithUnity ty => (x : ty) ->
x <.> inverse (unity {ty}) = inverse x
2 changes: 2 additions & 0 deletions libs/contrib/contrib.ipkg
Original file line number Diff line number Diff line change
@@ -1,7 +1,9 @@
package contrib

modules = Control.Delayed,

Control.Algebra,
Control.Algebra.Laws,

Data.List.TailRec,
Data.List.Equalities,
Expand Down

0 comments on commit 8eca37c

Please sign in to comment.