Files
agdaproofs/Numbers/Modulo/Addition.agda
2020-01-05 15:06:35 +00:00

50 lines
4.5 KiB
Agda
Raw Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

{-# OPTIONS --safe --warning=error --without-K #-}
-- These are explicitly with-K, because we currently encode an element of Zn as
-- a natural together with a proof that it is small.
open import LogicalFormulae
open import Numbers.Naturals.Semiring
open import Numbers.Naturals.Order
open import Semirings.Definition
open import Orders.Total.Definition
open import Numbers.Modulo.Definition
open import Numbers.Modulo.ModuloFunction
open import Numbers.Naturals.Order.Lemmas
module Numbers.Modulo.Addition where
open TotalOrder TotalOrder
_+n_ : {n : } .(pr : 0 <N n) n n pr n n pr n n pr
_+n_ {n} pr a b = record { x = mod n pr (n.x a +N n.x b) ; xLess = mod<N pr _ }
plusZnIdentityRight : {n : } .(pr : 0 <N n) (a : n n pr) (_+n_ pr a record { x = 0 ; xLess = pr }) a
plusZnIdentityRight 0<n record { x = x ; xLess = xLess } rewrite Semiring.sumZeroRight Semiring x = equalityZn (modIsMod 0<n x (<NProp xLess))
plusZnIdentityLeft : {n : } .(pr : 0 <N n) (a : n n pr) _+n_ pr (record { x = 0 ; xLess = pr }) a a
plusZnIdentityLeft 0<n record { x = x ; xLess = xLess } = equalityZn (modIsMod 0<n x (<NProp xLess))
plusZnCommutative : {n : } .(pr : 0 <N n) (a b : n n pr) _+n_ pr a b _+n_ pr b a
plusZnCommutative {n} 0<n a b = equalityZn (applyEquality (mod n 0<n) (Semiring.commutative Semiring (n.x a) _))
plusZnAssociative' : {n : } .(pr : 0 <N n) {a b c : } (a <N n) (b <N n) (c <N n) mod n pr ((mod n pr (a +N b)) +N c) mod n pr (a +N (mod n pr (b +N c)))
plusZnAssociative' 0<n {a} {b} {c} a<n b<n c<n with modAddition 0<n a<n b<n
plusZnAssociative' {n} 0<n {a} {b} {c} a<n b<n c<n | inl a+b=mod with modAddition 0<n b<n c<n
... | inl b+c=mod rewrite equalityCommutative (a+b=mod) | equalityCommutative (b+c=mod) = applyEquality (mod _ 0<n) (equalityCommutative (Semiring.+Associative Semiring a b c))
... | inr b+c=n+mod rewrite equalityCommutative (a+b=mod) | equalityCommutative (modAnd+n 0<n (a +N mod _ 0<n (b +N c))) | Semiring.+Associative Semiring n a (mod _ 0<n (b +N c)) | Semiring.commutative Semiring n a | equalityCommutative (Semiring.+Associative Semiring a n (mod _ 0<n (b +N c))) | equalityCommutative b+c=n+mod = applyEquality (mod _ 0<n) (equalityCommutative (Semiring.+Associative Semiring a b c))
plusZnAssociative' 0<n {a} {b} {c} a<n b<n c<n | inr a+b=n+mod with modAddition 0<n b<n c<n
plusZnAssociative' {n} 0<n {a} {b} {c} a<n b<n c<n | inr a+b=n+mod | inl b+c=mod rewrite equalityCommutative b+c=mod | equalityCommutative (modAnd+n 0<n (mod _ 0<n (a +N b) +N c)) | Semiring.+Associative Semiring n (mod _ 0<n (a +N b)) c | equalityCommutative a+b=n+mod = applyEquality (mod _ 0<n) (equalityCommutative (Semiring.+Associative Semiring a b c))
plusZnAssociative' {n} 0<n {a} {b} {c} a<n b<n c<n | inr a+b=n+mod | inr b+c=n+mod rewrite equalityCommutative (modAnd+n 0<n ((mod _ 0<n (a +N b)) +N c)) | equalityCommutative (modAnd+n 0<n (a +N mod _ 0<n (b +N c))) | Semiring.+Associative Semiring n (mod _ 0<n (a +N b)) c | equalityCommutative (a+b=n+mod) | Semiring.commutative Semiring a (mod n 0<n (b +N c)) | Semiring.+Associative Semiring n (mod n 0<n (b +N c)) a | equalityCommutative b+c=n+mod | Semiring.commutative Semiring (b +N c) a = applyEquality (mod _ 0<n) (equalityCommutative (Semiring.+Associative Semiring a b c))
plusZnAssociative : {n : } .(pr : 0 <N n) (a b c : n n pr) _+n_ pr (_+n_ pr a b) c _+n_ pr a (_+n_ pr b c)
plusZnAssociative {succ n} 0<sn record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } = equalityZn t
where
ma = mod (succ n) 0<sn a
mb = mod (succ n) 0<sn b
mc = mod (succ n) 0<sn c
v : mod (succ n) 0<sn ((mod (succ n) 0<sn (ma +N mb)) +N mc) mod (succ n) 0<sn (ma +N mod (succ n) 0<sn (mb +N mc))
v = plusZnAssociative' 0<sn (mod<N 0<sn a) (mod<N 0<sn b) (mod<N 0<sn c)
u : mod (succ n) 0<sn ((mod (succ n) 0<sn (mod (succ n) 0<sn a +N (mod (succ n) 0<sn b))) +N c) mod (succ n) 0<sn (a +N mod (succ n) 0<sn ((mod (succ n) 0<sn b) +N (mod (succ n) 0<sn c)))
u rewrite equalityCommutative (modIsMod 0<sn c (<NProp cLess)) | equalityCommutative (modIsMod 0<sn a (<NProp aLess)) | modIdempotent 0<sn a | modIdempotent 0<sn c = v
t : mod (succ n) 0<sn (mod (succ n) 0<sn (a +N b) +N c) mod (succ n) 0<sn (a +N mod (succ n) 0<sn (b +N c))
t rewrite modExtracts {succ n} 0<sn a b | modExtracts {succ n} 0<sn b c = u