|
| 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 | +-- selfSquareId x p = |
| 29 | +-- rewrite sym $ monoidNeutralIsNeutralR x in |
| 30 | +-- rewrite sym $ groupInverseIsInverseR x in |
| 31 | +-- rewrite sym $ semigroupOpIsAssociative (inverse x) x x in |
| 32 | +-- rewrite p in |
| 33 | +-- Refl |
| 34 | + |
| 35 | +||| Inverse elements commute. |
| 36 | +public export |
| 37 | +inverseCommute : VGroup ty => (x, y : ty) -> |
| 38 | + y <+> x = neutral {ty} -> x <+> y = neutral {ty} |
| 39 | +-- inverseCommute x y p = selfSquareId (x <+> y) prop where |
| 40 | +-- prop : (x <+> y) <+> (x <+> y) = x <+> y |
| 41 | +-- prop = |
| 42 | +-- rewrite sym $ semigroupOpIsAssociative x y (x <+> y) in |
| 43 | +-- rewrite semigroupOpIsAssociative y x y in |
| 44 | +-- rewrite p in |
| 45 | +-- rewrite monoidNeutralIsNeutralR y in |
| 46 | +-- Refl |
| 47 | + |
| 48 | +||| Every element has a right inverse. |
| 49 | +public export |
| 50 | +groupInverseIsInverseL : VGroup ty => (x : ty) -> |
| 51 | + x <+> inverse x = neutral {ty} |
| 52 | +-- groupInverseIsInverseL x = |
| 53 | +-- inverseCommute x (inverse x) (groupInverseIsInverseR x) |
| 54 | + |
| 55 | +||| -(-x) = x |
| 56 | +public export |
| 57 | +inverseSquaredIsIdentity : VGroup ty => (x : ty) -> |
| 58 | + inverse (inverse x) = x |
| 59 | +-- inverseSquaredIsIdentity x = |
| 60 | +-- let x' = inverse x in |
| 61 | +-- uniqueInverse |
| 62 | +-- x' |
| 63 | +-- (inverse x') |
| 64 | +-- x |
| 65 | +-- (groupInverseIsInverseR x') |
| 66 | +-- (groupInverseIsInverseR x) |
| 67 | + |
| 68 | +||| If every square in a group is identity, the group is commutative. |
| 69 | +public export |
| 70 | +squareIdCommutative : VGroup ty => (x, y : ty) -> |
| 71 | + ((a : ty) -> a <+> a = neutral {ty}) -> |
| 72 | + x <+> y = y <+> x |
| 73 | +-- squareIdCommutative x y p = |
| 74 | +-- let |
| 75 | +-- xy = x <+> y |
| 76 | +-- yx = y <+> x |
| 77 | +-- in |
| 78 | +-- uniqueInverse xy xy yx (p xy) prop where |
| 79 | +-- prop : (x <+> y) <+> (y <+> x) = A.neutral |
| 80 | +-- prop = |
| 81 | +-- rewrite sym $ semigroupOpIsAssociative x y (y <+> x) in |
| 82 | +-- rewrite semigroupOpIsAssociative y y x in |
| 83 | +-- rewrite p y in |
| 84 | +-- rewrite monoidNeutralIsNeutralR x in |
| 85 | +-- p x |
| 86 | + |
| 87 | +||| -0 = 0 |
| 88 | +public export |
| 89 | +inverseNeutralIsNeutral : VGroup ty => |
| 90 | + inverse (neutral {ty}) = neutral {ty} |
| 91 | +-- inverseNeutralIsNeutral {ty} = |
| 92 | +-- let e = neutral {ty} in |
| 93 | +-- rewrite sym $ cong {f = inverse} (groupInverseIsInverseL e) in |
| 94 | +-- rewrite monoidNeutralIsNeutralR $ inverse e in |
| 95 | +-- inverseSquaredIsIdentity e |
| 96 | + |
| 97 | +||| -(x + y) = -y + -x |
| 98 | +public export |
| 99 | +inverseOfSum : VGroup ty => (l, r : ty) -> |
| 100 | + inverse (l <+> r) = inverse r <+> inverse l |
| 101 | +-- inverseOfSum {ty} l r = |
| 102 | +-- let |
| 103 | +-- e = neutral {ty} |
| 104 | +-- il = inverse l |
| 105 | +-- ir = inverse r |
| 106 | +-- lr = l <+> r |
| 107 | +-- ilr = inverse lr |
| 108 | +-- iril = ir <+> il |
| 109 | +-- ile = il <+> e |
| 110 | +-- in |
| 111 | +-- -- expand |
| 112 | +-- rewrite sym $ monoidNeutralIsNeutralR ilr in |
| 113 | +-- rewrite sym $ groupInverseIsInverseR r in |
| 114 | +-- rewrite sym $ monoidNeutralIsNeutralL ir in |
| 115 | +-- rewrite sym $ groupInverseIsInverseR l in |
| 116 | +-- -- shuffle |
| 117 | +-- rewrite semigroupOpIsAssociative ir il l in |
| 118 | +-- rewrite sym $ semigroupOpIsAssociative iril l r in |
| 119 | +-- rewrite sym $ semigroupOpIsAssociative iril lr ilr in |
| 120 | +-- -- contract |
| 121 | +-- rewrite sym $ monoidNeutralIsNeutralL il in |
| 122 | +-- rewrite groupInverseIsInverseL lr in |
| 123 | +-- rewrite sym $ semigroupOpIsAssociative (ir <+> ile) l ile in |
| 124 | +-- rewrite semigroupOpIsAssociative l il e in |
| 125 | +-- rewrite groupInverseIsInverseL l in |
| 126 | +-- rewrite monoidNeutralIsNeutralL e in |
| 127 | +-- Refl |
| 128 | + |
| 129 | +||| y = z if x + y = x + z |
| 130 | +public export |
| 131 | +cancelLeft : VGroup ty => (x, y, z : ty) -> |
| 132 | + x <+> y = x <+> z -> y = z |
| 133 | +cancelLeft x y z p = |
| 134 | + rewrite sym $ monoidNeutralIsNeutralR y in |
| 135 | + rewrite sym $ groupInverseIsInverseR x in |
| 136 | + rewrite sym $ semigroupOpIsAssociative (inverse x) x y in |
| 137 | + rewrite p in |
| 138 | + rewrite semigroupOpIsAssociative (inverse x) x z in |
| 139 | + rewrite groupInverseIsInverseR x in |
| 140 | + monoidNeutralIsNeutralR z |
| 141 | + |
| 142 | +||| y = z if y + x = z + x. |
| 143 | +public export |
| 144 | +cancelRight : VGroup ty => (x, y, z : ty) -> |
| 145 | + y <+> x = z <+> x -> y = z |
| 146 | +-- cancelRight x y z p = |
| 147 | +-- rewrite sym $ monoidNeutralIsNeutralL y in |
| 148 | +-- rewrite sym $ groupInverseIsInverseL x in |
| 149 | +-- rewrite semigroupOpIsAssociative y x (inverse x) in |
| 150 | +-- rewrite p in |
| 151 | +-- rewrite sym $ semigroupOpIsAssociative z x (inverse x) in |
| 152 | +-- rewrite groupInverseIsInverseL x in |
| 153 | +-- monoidNeutralIsNeutralL z |
| 154 | + |
| 155 | +||| ab = 0 -> a = b^-1 |
| 156 | +public export |
| 157 | +neutralProductInverseR : VGroup ty => (a, b : ty) -> |
| 158 | + a <+> b = neutral {ty} -> a = inverse b |
| 159 | +-- neutralProductInverseR a b prf = |
| 160 | +-- cancelRight b a (inverse b) $ |
| 161 | +-- trans prf $ sym $ groupInverseIsInverseR b |
| 162 | + |
| 163 | +||| ab = 0 -> a^-1 = b |
| 164 | +public export |
| 165 | +neutralProductInverseL : VGroup ty => (a, b : ty) -> |
| 166 | + a <+> b = neutral {ty} -> inverse a = b |
| 167 | +neutralProductInverseL a b prf = |
| 168 | + cancelLeft a (inverse a) b $ |
| 169 | + trans (groupInverseIsInverseL a) $ sym prf |
| 170 | + |
| 171 | +||| For any a and b, ax = b and ya = b have solutions. |
| 172 | +public export |
| 173 | +latinSquareProperty : VGroup ty => (a, b : ty) -> |
| 174 | + ((x : ty ** a <+> x = b), |
| 175 | + (y : ty ** y <+> a = b)) |
| 176 | +-- latinSquareProperty a b = |
| 177 | +-- let a' = inverse a in |
| 178 | +-- (((a' <+> b) ** |
| 179 | +-- rewrite semigroupOpIsAssociative a a' b in |
| 180 | +-- rewrite groupInverseIsInverseL a in |
| 181 | +-- monoidNeutralIsNeutralR b), |
| 182 | +-- (b <+> a' ** |
| 183 | +-- rewrite sym $ semigroupOpIsAssociative b a' a in |
| 184 | +-- rewrite groupInverseIsInverseR a in |
| 185 | +-- monoidNeutralIsNeutralL b)) |
| 186 | + |
| 187 | +||| For any a, b, x, the solution to ax = b is unique. |
| 188 | +public export |
| 189 | +uniqueSolutionR : VGroup ty => (a, b, x, y : ty) -> |
| 190 | + a <+> x = b -> a <+> y = b -> x = y |
| 191 | +uniqueSolutionR a b x y p q = cancelLeft a x y $ trans p (sym q) |
| 192 | + |
| 193 | +||| For any a, b, y, the solution to ya = b is unique. |
| 194 | +public export |
| 195 | +uniqueSolutionL : VGroup t => (a, b, x, y : t) -> |
| 196 | + x <+> a = b -> y <+> a = b -> x = y |
| 197 | +uniqueSolutionL a b x y p q = cancelRight a x y $ trans p (sym q) |
| 198 | + |
| 199 | +||| -(x + y) = -x + -y |
| 200 | +public export |
| 201 | +inverseDistributesOverGroupOp : AbelianGroup ty => (l, r : ty) -> |
| 202 | + inverse (l <+> r) = inverse l <+> inverse r |
| 203 | +inverseDistributesOverGroupOp l r = |
| 204 | + rewrite groupOpIsCommutative (inverse l) (inverse r) in |
| 205 | + inverseOfSum l r |
| 206 | + |
| 207 | +||| Homomorphism preserves neutral. |
| 208 | +public export |
| 209 | +homoNeutral : GroupHomomorphism a b => |
| 210 | + to (neutral {ty=a}) = neutral {ty=b} |
| 211 | +-- homoNeutral {a} = |
| 212 | +-- let ae = neutral {ty=a} in |
| 213 | +-- selfSquareId (to ae) $ |
| 214 | +-- trans (sym $ toGroup ae ae) $ |
| 215 | +-- cong {f = to} $ monoidNeutralIsNeutralL ae |
| 216 | + |
| 217 | +||| Homomorphism preserves inverses. |
| 218 | +public export |
| 219 | +homoInverse : GroupHomomorphism a b => (x : a) -> |
| 220 | + the b (to $ inverse x) = the b (inverse $ to x) |
| 221 | +-- homoInverse {a} {b} x = |
| 222 | +-- cancelRight (to x) (to $ inverse x) (inverse $ to x) $ |
| 223 | +-- trans (sym $ toGroup (inverse x) x) $ |
| 224 | +-- trans (cong {f = to} $ groupInverseIsInverseR x) $ |
| 225 | +-- trans (homoNeutral {a=a} {b=b}) $ |
| 226 | +-- sym $ groupInverseIsInverseR (to x) |
| 227 | + |
| 228 | +-- Rings |
| 229 | + |
| 230 | +||| 0x = x |
| 231 | +public export |
| 232 | +multNeutralAbsorbingL : VRing ty => (r : ty) -> |
| 233 | + neutral {ty} <.> r = neutral {ty} |
| 234 | +-- multNeutralAbsorbingL {ty} r = |
| 235 | +-- let |
| 236 | +-- e = neutral {ty} |
| 237 | +-- ir = inverse r |
| 238 | +-- exr = e <.> r |
| 239 | +-- iexr = inverse exr |
| 240 | +-- in |
| 241 | +-- rewrite sym $ monoidNeutralIsNeutralR exr in |
| 242 | +-- rewrite sym $ groupInverseIsInverseR exr in |
| 243 | +-- rewrite sym $ semigroupOpIsAssociative iexr exr ((iexr <+> exr) <.> r) in |
| 244 | +-- rewrite groupInverseIsInverseR exr in |
| 245 | +-- rewrite sym $ ringOpIsDistributiveR e e r in |
| 246 | +-- rewrite monoidNeutralIsNeutralR e in |
| 247 | +-- groupInverseIsInverseR exr |
| 248 | + |
| 249 | +||| x0 = 0 |
| 250 | +public export |
| 251 | +multNeutralAbsorbingR : VRing ty => (l : ty) -> |
| 252 | + l <.> neutral {ty} = neutral {ty} |
| 253 | +-- multNeutralAbsorbingR {ty} l = |
| 254 | +-- let |
| 255 | +-- e = neutral {ty} |
| 256 | +-- il = inverse l |
| 257 | +-- lxe = l <.> e |
| 258 | +-- ilxe = inverse lxe |
| 259 | +-- in |
| 260 | +-- rewrite sym $ monoidNeutralIsNeutralL lxe in |
| 261 | +-- rewrite sym $ groupInverseIsInverseL lxe in |
| 262 | +-- rewrite semigroupOpIsAssociative (l <.> (lxe <+> ilxe)) lxe ilxe in |
| 263 | +-- rewrite groupInverseIsInverseL lxe in |
| 264 | +-- rewrite sym $ ringOpIsDistributiveL l e e in |
| 265 | +-- rewrite monoidNeutralIsNeutralL e in |
| 266 | +-- groupInverseIsInverseL lxe |
| 267 | + |
| 268 | +||| (-x)y = -(xy) |
| 269 | +public export |
| 270 | +multInverseInversesL : VRing ty => (l, r : ty) -> |
| 271 | + inverse l <.> r = inverse (l <.> r) |
| 272 | +-- multInverseInversesL l r = |
| 273 | +-- let |
| 274 | +-- il = inverse l |
| 275 | +-- lxr = l <.> r |
| 276 | +-- ilxr = il <.> r |
| 277 | +-- i_lxr = inverse lxr |
| 278 | +-- in |
| 279 | +-- rewrite sym $ monoidNeutralIsNeutralR ilxr in |
| 280 | +-- rewrite sym $ groupInverseIsInverseR lxr in |
| 281 | +-- rewrite sym $ semigroupOpIsAssociative i_lxr lxr ilxr in |
| 282 | +-- rewrite sym $ ringOpIsDistributiveR l il r in |
| 283 | +-- rewrite groupInverseIsInverseL l in |
| 284 | +-- rewrite multNeutralAbsorbingL r in |
| 285 | +-- monoidNeutralIsNeutralL i_lxr |
| 286 | + |
| 287 | +||| x(-y) = -(xy) |
| 288 | +public export |
| 289 | +multInverseInversesR : VRing ty => (l, r : ty) -> |
| 290 | + l <.> inverse r = inverse (l <.> r) |
| 291 | +-- multInverseInversesR l r = |
| 292 | +-- let |
| 293 | +-- ir = inverse r |
| 294 | +-- lxr = l <.> r |
| 295 | +-- lxir = l <.> ir |
| 296 | +-- ilxr = inverse lxr |
| 297 | +-- in |
| 298 | +-- rewrite sym $ monoidNeutralIsNeutralL lxir in |
| 299 | +-- rewrite sym $ groupInverseIsInverseL lxr in |
| 300 | +-- rewrite semigroupOpIsAssociative lxir lxr ilxr in |
| 301 | +-- rewrite sym $ ringOpIsDistributiveL l ir r in |
| 302 | +-- rewrite groupInverseIsInverseR r in |
| 303 | +-- rewrite multNeutralAbsorbingR l in |
| 304 | +-- monoidNeutralIsNeutralR ilxr |
| 305 | + |
| 306 | +||| (-x)(-y) = xy |
| 307 | +public export |
| 308 | +multNegativeByNegativeIsPositive : VRing ty => (l, r : ty) -> |
| 309 | + inverse l <.> inverse r = l <.> r |
| 310 | +-- multNegativeByNegativeIsPositive l r = |
| 311 | +-- rewrite multInverseInversesR (inverse l) r in |
| 312 | +-- rewrite sym $ multInverseInversesL (inverse l) r in |
| 313 | +-- rewrite inverseSquaredIsIdentity l in |
| 314 | +-- Refl |
| 315 | + |
| 316 | +||| (-1)x = -x |
| 317 | +public export |
| 318 | +inverseOfUnityR : VRingWithUnity ty => (x : ty) -> |
| 319 | + inverse (unity {ty}) <.> x = inverse x |
| 320 | +-- inverseOfUnityR {ty} x = |
| 321 | +-- rewrite multInverseInversesL (unity {ty}) x in |
| 322 | +-- rewrite ringWithUnityIsUnityR x in |
| 323 | +-- Refl |
| 324 | + |
| 325 | +||| x(-1) = -x |
| 326 | +public export |
| 327 | +inverseOfUnityL : VRingWithUnity ty => (x : ty) -> |
| 328 | + x <.> inverse (unity {ty}) = inverse x |
| 329 | +-- inverseOfUnityL {ty} x = |
| 330 | +-- rewrite multInverseInversesR x (unity {ty}) in |
| 331 | +-- rewrite ringWithUnityIsUnityL x in |
| 332 | +-- Refl |
0 commit comments