mirror of
https://github.com/Smaug123/agdaproofs
synced 2025-10-18 09:48:38 +00:00
Cleanup finset and modulo (#92)
This commit is contained in:
@@ -1,4 +1,4 @@
|
||||
{-# OPTIONS --safe --warning=error #-}
|
||||
{-# 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.
|
||||
|
||||
@@ -10,95 +10,48 @@ open import Groups.FiniteGroups.Definition
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Order
|
||||
open import Numbers.Primes.PrimeNumbers
|
||||
open import Setoids.Setoids
|
||||
open import Sets.FinSet
|
||||
open import Sets.FinSetWithK
|
||||
open import Sets.FinSet.Definition
|
||||
open import Functions
|
||||
open import Numbers.Naturals.WithK
|
||||
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
|
||||
|
||||
cancelSumFromInequality : {a b c : ℕ} → a +N b <N a +N c → b <N c
|
||||
cancelSumFromInequality {a} {b} {c} (le x proof) = le x help
|
||||
where
|
||||
help : succ x +N b ≡ c
|
||||
help rewrite Semiring.+Associative ℕSemiring (succ x) a b | Semiring.commutative ℕSemiring x a | equalityCommutative (Semiring.+Associative ℕSemiring (succ a) x b) | Semiring.commutative ℕSemiring a (x +N b) | Semiring.commutative ℕSemiring (succ (x +N b)) a = canSubtractFromEqualityLeft {a} proof
|
||||
_+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 _ }
|
||||
|
||||
_+n_ : {n : ℕ} {pr : 0 <N n} → ℤn n pr → ℤn n pr → ℤn n pr
|
||||
_+n_ {0} {le x ()} a b
|
||||
_+n_ {succ n} {pr} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } with totality (a +N b) (succ n)
|
||||
_+n_ {succ n} {pr} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inl (inl (a+b<n)) = record { x = a +N b ; xLess = a+b<n }
|
||||
_+n_ {succ n} {pr} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inl (inr (n<a+b)) = record { x = subtractionNResult.result sub ; xLess = pr2 }
|
||||
where
|
||||
sub : subtractionNResult (succ n) (a +N b) (inl n<a+b)
|
||||
sub = -N (inl n<a+b)
|
||||
pr1 : a +N b <N succ n +N succ n
|
||||
pr1 = addStrongInequalities aLess bLess
|
||||
pr1' : succ n +N (subtractionNResult.result sub) <N succ n +N succ n
|
||||
pr1' rewrite subtractionNResult.pr sub = pr1
|
||||
pr2 : subtractionNResult.result sub <N succ n
|
||||
pr2 = cancelSumFromInequality pr1'
|
||||
_+n_ {succ n} {pr} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inr (a+b=n) = record { x = 0 ; xLess = succIsPositive n }
|
||||
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))
|
||||
|
||||
plusZnIdentityRight : {n : ℕ} → {pr : 0 <N n} → (a : ℤn n pr) → (a +n record { x = 0 ; xLess = pr }) ≡ a
|
||||
plusZnIdentityRight {zero} {()} a
|
||||
plusZnIdentityRight {succ x} {_} record { x = a ; xLess = aLess } with totality (a +N 0) (succ x)
|
||||
plusZnIdentityRight {succ x} {_} record { x = a ; xLess = aLess } | inl (inl a<sx) = equalityZn _ _ (Semiring.commutative ℕSemiring a 0)
|
||||
plusZnIdentityRight {succ x} {_} record { x = a ; xLess = aLess } | inl (inr sx<a) = exFalso (f aLess sx<a)
|
||||
where
|
||||
f : (aL : a <N succ x) → (sx<a : succ x <N a +N 0) → False
|
||||
f aL sx<a rewrite Semiring.commutative ℕSemiring a 0 = irreflexive (<Transitive aL sx<a)
|
||||
plusZnIdentityRight {succ x} {_} record { x = a ; xLess = aLess } | inr a=sx = exFalso (f aLess a=sx)
|
||||
where
|
||||
f : (aL : a <N succ x) → (a=sx : a +N 0 ≡ succ x) → False
|
||||
f aL a=sx rewrite Semiring.commutative ℕSemiring a 0 | a=sx = TotalOrder.irreflexive ℕTotalOrder aL
|
||||
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) _))
|
||||
|
||||
plusZnIdentityLeft : {n : ℕ} → {pr : 0 <N n} → (a : ℤn n pr) → (record { x = 0 ; xLess = pr }) +n a ≡ a
|
||||
plusZnIdentityLeft {zero} {()}
|
||||
plusZnIdentityLeft {succ n} {pr} record { x = x ; xLess = xLess } with totality x (succ n)
|
||||
plusZnIdentityLeft {succ n} {pr} record { x = x ; xLess = xLess } | inl (inl x<succn) rewrite <NRefl x<succn xLess = refl
|
||||
plusZnIdentityLeft {succ n} {pr} record { x = x ; xLess = xLess } | inl (inr succn<x) = exFalso (TotalOrder.irreflexive ℕTotalOrder (TotalOrder.<Transitive ℕTotalOrder succn<x xLess))
|
||||
plusZnIdentityLeft {succ n} {pr} record { x = x ; xLess = xLess } | inr x=succn rewrite x=succn = exFalso (TotalOrder.irreflexive ℕTotalOrder xLess)
|
||||
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))
|
||||
|
||||
subLemma : {a b c : ℕ} → (pr1 : a <N b) → (pr2 : a <N c) → (pr : b ≡ c) → (subtractionNResult.result (-N (inl pr1))) ≡ (subtractionNResult.result (-N (inl pr2)))
|
||||
subLemma {a} {b} {c} a<b a<c b=c rewrite b=c | <NRefl a<b a<c = refl
|
||||
|
||||
plusZnCommutative : {n : ℕ} → {pr : 0 <N n} → (a b : ℤn n pr) → (a +n b) ≡ b +n a
|
||||
plusZnCommutative {zero} {()} a b
|
||||
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } with totality (a +N b) (succ n)
|
||||
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inl (inl a+b<sn) with totality (b +N a) (succ n)
|
||||
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inl (inl a+b<sn) | inl (inl b+a<sn) = equalityZn _ _ (Semiring.commutative ℕSemiring a b)
|
||||
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inl (inl a+b<sn) | inl (inr sn<b+a) = exFalso (f a+b<sn sn<b+a)
|
||||
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
|
||||
f : (a +N b) <N succ n → succ n <N b +N a → False
|
||||
f pr1 pr2 rewrite Semiring.commutative ℕSemiring b a = TotalOrder.irreflexive ℕTotalOrder (<Transitive pr2 pr1)
|
||||
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inl (inl a+b<sn) | inr b+a=sn = exFalso (f a+b<sn b+a=sn)
|
||||
where
|
||||
f : (a +N b) <N succ n → b +N a ≡ succ n → False
|
||||
f pr1 eq rewrite Semiring.commutative ℕSemiring b a | eq = TotalOrder.irreflexive ℕTotalOrder pr1
|
||||
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inl (inr sn<a+b) with totality (b +N a) (succ n)
|
||||
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inl (inr sn<a+b) | inl (inl b+a<sn) = exFalso (f sn<a+b b+a<sn)
|
||||
where
|
||||
f : succ n <N a +N b → b +N a <N succ n → False
|
||||
f pr1 pr2 rewrite Semiring.commutative ℕSemiring b a = TotalOrder.irreflexive ℕTotalOrder (<Transitive sn<a+b b+a<sn)
|
||||
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inl (inr sn<a+b) | inl (inr sn<b+a) = equalityZn _ _ (subLemma sn<a+b sn<b+a (Semiring.commutative ℕSemiring a b))
|
||||
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inl (inr sn<a+b) | inr sn=b+a = exFalso (f sn<a+b sn=b+a)
|
||||
where
|
||||
f : succ n <N a +N b → b +N a ≡ succ n → False
|
||||
f pr eq rewrite Semiring.commutative ℕSemiring b a | eq = TotalOrder.irreflexive ℕTotalOrder pr
|
||||
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inr sn=a+b with totality (b +N a) (succ n)
|
||||
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inr sn=a+b | inl (inl b+a<sn) = exFalso f
|
||||
where
|
||||
f : False
|
||||
f rewrite Semiring.commutative ℕSemiring b a | sn=a+b = TotalOrder.irreflexive ℕTotalOrder b+a<sn
|
||||
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inr sn=a+b | inl (inr sn<b+a) = exFalso f
|
||||
where
|
||||
f : False
|
||||
f rewrite Semiring.commutative ℕSemiring b a | sn=a+b = TotalOrder.irreflexive ℕTotalOrder sn<b+a
|
||||
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inr sn=a+b | inr sn=b+a = equalityZn _ _ refl
|
||||
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
|
||||
|
@@ -1,6 +1,4 @@
|
||||
{-# OPTIONS --safe --warning=error #-}
|
||||
-- These are explicitly with-K, because we currently encode an element of Zn as
|
||||
-- a natural together with a proof that it is small.
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Groups.Definition
|
||||
@@ -8,20 +6,22 @@ open import Groups.Groups
|
||||
open import Groups.Abelian.Definition
|
||||
open import Groups.FiniteGroups.Definition
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Order
|
||||
open import Numbers.Primes.PrimeNumbers
|
||||
open import Setoids.Setoids
|
||||
open import Sets.FinSet
|
||||
open import Sets.FinSetWithK
|
||||
open import Sets.FinSet.Definition
|
||||
open import Functions
|
||||
open import Numbers.Naturals.WithK
|
||||
open import Semirings.Definition
|
||||
open import Numbers.Naturals.EuclideanAlgorithm
|
||||
open import Orders.Total.Definition
|
||||
|
||||
module Numbers.Modulo.Definition where
|
||||
record ℤn (n : ℕ) (pr : 0 <N n) : Set where
|
||||
field
|
||||
x : ℕ
|
||||
xLess : x <N n
|
||||
|
||||
equalityZn : {n : ℕ} {pr : 0 <N n} → (a b : ℤn n pr) → (ℤn.x a ≡ ℤn.x b) → a ≡ b
|
||||
equalityZn record { x = a ; xLess = aLess } record { x = .a ; xLess = bLess } refl rewrite <NRefl aLess bLess = refl
|
||||
record ℤn (n : ℕ) .(pr : 0 <N n) : Set where
|
||||
field
|
||||
x : ℕ
|
||||
.xLess : x <N n
|
||||
|
||||
equalityZn : {n : ℕ} .{pr : 0 <N n} → {a b : ℤn n pr} → (ℤn.x a ≡ ℤn.x b) → a ≡ b
|
||||
equalityZn {a = record { x = a ; xLess = aLess }} {record { x = .a ; xLess = bLess }} refl = refl
|
||||
|
||||
|
@@ -1,6 +1,4 @@
|
||||
{-# OPTIONS --safe --warning=error #-}
|
||||
-- These are explicitly with-K, because we currently encode an element of Zn as
|
||||
-- a natural together with a proof that it is small.
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Groups.Definition
|
||||
@@ -11,616 +9,53 @@ open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
open import Numbers.Naturals.Order.Lemmas
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Primes.PrimeNumbers
|
||||
open import Setoids.Setoids
|
||||
open import Sets.FinSet
|
||||
open import Sets.FinSetWithK
|
||||
open import Sets.FinSet.Definition
|
||||
open import Sets.FinSet.Lemmas
|
||||
open import Sets.Cardinality.Finite.Definition
|
||||
open import Functions
|
||||
open import Numbers.Naturals.WithK
|
||||
open import Semirings.Definition
|
||||
open import Numbers.Modulo.Definition
|
||||
open import Numbers.Modulo.Addition
|
||||
open import Orders.Total.Definition
|
||||
open import Numbers.Naturals.EuclideanAlgorithm
|
||||
open import Numbers.Modulo.ModuloFunction
|
||||
|
||||
module Numbers.Modulo.Group where
|
||||
|
||||
open TotalOrder ℕTotalOrder
|
||||
open Semiring ℕSemiring
|
||||
|
||||
help30 : {a b c n : ℕ} → (c <N n) → (a +N b ≡ n) → (n<b+c : n <N b +N c) → (n <N a +N subtractionNResult.result (-N (inl n<b+c))) → False
|
||||
help30 {a} {b} {c} {n} c<n a+b=n n<b+c x = TotalOrder.irreflexive ℕTotalOrder (<Transitive pr5 c<n)
|
||||
where
|
||||
pr : n +N n <N a +N (subtractionNResult.result (-N (inl n<b+c)) +N n)
|
||||
pr = identityOfIndiscernablesRight _<N_ (additionPreservesInequality n x) (equalityCommutative (Semiring.+Associative ℕSemiring a _ n))
|
||||
pr2 : n +N n <N a +N (b +N c)
|
||||
pr2 = identityOfIndiscernablesRight _<N_ pr (applyEquality (a +N_) (addMinus (inl n<b+c)))
|
||||
pr3 : n +N n <N (a +N b) +N c
|
||||
pr3 rewrite equalityCommutative (Semiring.+Associative ℕSemiring a b c) = pr2
|
||||
pr4 : n +N n <N c +N n
|
||||
pr4 rewrite Semiring.commutative ℕSemiring c n = identityOfIndiscernablesRight _<N_ pr3 (applyEquality (_+N c) a+b=n)
|
||||
pr5 : n <N c
|
||||
pr5 = subtractionPreservesInequality n pr4
|
||||
private
|
||||
0<s : {n : ℕ} → 0 <N succ n
|
||||
0<s {n} = le n (applyEquality succ (Semiring.sumZeroRight ℕSemiring n))
|
||||
|
||||
help31 : {a b c n : ℕ} → (a +N b ≡ n) → (n<b+c : n <N b +N c) → (a +N subtractionNResult.result (-N (inl n<b+c))) ≡ c
|
||||
help31 {a} {b} {c} {n} a+b=n n<b+c = canSubtractFromEqualityLeft pr2
|
||||
where
|
||||
pr1 : a +N (n +N subtractionNResult.result (-N (inl n<b+c))) ≡ n +N c
|
||||
pr1 rewrite addMinus' (inl n<b+c) | Semiring.+Associative ℕSemiring a b c | a+b=n = refl
|
||||
pr2 : n +N (a +N subtractionNResult.result (-N (inl n<b+c))) ≡ n +N c
|
||||
pr2 = identityOfIndiscernablesLeft _≡_ pr1 (lemm a n (subtractionNResult.result (-N (inl n<b+c))))
|
||||
where
|
||||
lemm : (a b c : ℕ) → a +N (b +N c) ≡ b +N (a +N c)
|
||||
lemm a b c rewrite Semiring.+Associative ℕSemiring a b c | Semiring.commutative ℕSemiring a b | equalityCommutative (Semiring.+Associative ℕSemiring b a c) = refl
|
||||
inverseN : {n : ℕ} → .(0<n : 0 <N n) → (x : ℤn n 0<n) → ℤn n 0<n
|
||||
inverseN 0<n record { x = 0 ; xLess = _ } = record { x = 0 ; xLess = 0<n }
|
||||
inverseN 0<n record { x = succ x ; xLess = xLess } with <NProp xLess
|
||||
... | le subtr pr = record { x = succ subtr ; xLess = le x (transitivity (commutative (succ x) (succ subtr)) pr) }
|
||||
|
||||
help7 : {a b c n : ℕ} → b +N c ≡ n → a <N n → (n<a+b : n <N a +N b) → (subtractionNResult.result (-N (inl n<a+b)) +N c ≡ n) → False
|
||||
help7 {a} {b} {c} {n} b+c=n a<n n<a+b x = TotalOrder.irreflexive ℕTotalOrder (identityOfIndiscernablesLeft _<N_ a<n pr5)
|
||||
where
|
||||
pr : (n +N subtractionNResult.result (-N (inl n<a+b))) +N c ≡ n +N n
|
||||
pr = identityOfIndiscernablesLeft _≡_ (applyEquality (n +N_) x) (Semiring.+Associative ℕSemiring n _ c)
|
||||
pr2 : (a +N b) +N c ≡ n +N n
|
||||
pr2 = identityOfIndiscernablesLeft _≡_ pr (applyEquality (_+N c) (addMinus' (inl n<a+b)))
|
||||
pr3 : a +N (b +N c) ≡ n +N n
|
||||
pr3 rewrite Semiring.+Associative ℕSemiring a b c = pr2
|
||||
pr4 : a +N n ≡ n +N n
|
||||
pr4 = identityOfIndiscernablesLeft _≡_ pr3 (applyEquality (a +N_) b+c=n)
|
||||
pr5 : a ≡ n
|
||||
pr5 = canSubtractFromEqualityRight pr4
|
||||
invLeft : {n : ℕ} → .(0<n : 0 <N n) → (x : ℤn n 0<n) → _+n_ 0<n (inverseN 0<n x) x ≡ record { x = 0 ; xLess = 0<n }
|
||||
invLeft {n} 0<n record { x = 0 ; xLess = xLess } = plusZnIdentityLeft 0<n (record { x = 0 ; xLess = 0<n })
|
||||
invLeft {n} 0<n record { x = (succ x) ; xLess = xLess } with <NProp xLess
|
||||
... | le subtr pr rewrite pr = equalityZn (modN 0<n)
|
||||
|
||||
help29 : {a b c n : ℕ} → (c <N n) → (n<b+c : n <N b +N c) → (a +N subtractionNResult.result (-N (inl n<b+c))) ≡ n → a +N b ≡ n → False
|
||||
help29 {a} {b} {c} {n} c<n n<b+c pr a+b=n = TotalOrder.irreflexive ℕTotalOrder (identityOfIndiscernablesLeft _<N_ c<n p4)
|
||||
where
|
||||
p1 : a +N (subtractionNResult.result (-N (inl n<b+c)) +N n) ≡ n +N n
|
||||
p1 = identityOfIndiscernablesLeft _≡_ (applyEquality (_+N n) pr) (equalityCommutative (Semiring.+Associative ℕSemiring a _ n))
|
||||
p2 : (a +N b) +N c ≡ n +N n
|
||||
p2 rewrite equalityCommutative (Semiring.+Associative ℕSemiring a b c) = identityOfIndiscernablesLeft _≡_ p1 (applyEquality (a +N_) (addMinus (inl n<b+c)))
|
||||
p3 : n +N c ≡ n +N n
|
||||
p3 = identityOfIndiscernablesLeft _≡_ p2 (applyEquality (_+N c) a+b=n)
|
||||
p4 : c ≡ n
|
||||
p4 = canSubtractFromEqualityLeft p3
|
||||
|
||||
help28 : {a b c n : ℕ} → (n<a+'b+c : n <N a +N (b +N c)) → (a +N b ≡ n) → subtractionNResult.result (-N (inl n<a+'b+c)) ≡ c
|
||||
help28 {a} {b} {c} {n} pr a+b=n = canSubtractFromEqualityLeft p2
|
||||
where
|
||||
p1 : a +N (b +N c) ≡ n +N c
|
||||
p1 rewrite Semiring.+Associative ℕSemiring a b c | a+b=n = refl
|
||||
p2 : n +N subtractionNResult.result (-N (inl pr)) ≡ n +N c
|
||||
p2 = identityOfIndiscernablesLeft _≡_ p1 (equalityCommutative (addMinus' (inl pr)))
|
||||
|
||||
help27 : {a b c n : ℕ} → (a +N b ≡ succ n) → (a +N (b +N c) <N succ n) → a +N (b +N c) ≡ c
|
||||
help27 {a} {b} {zero} {n} a+b=sn a+b+c<sn rewrite Semiring.commutative ℕSemiring b 0 | a+b=sn = exFalso (TotalOrder.irreflexive ℕTotalOrder a+b+c<sn)
|
||||
help27 {a} {b} {succ c} {n} a+b=sn a+b+c<sn rewrite Semiring.+Associative ℕSemiring a b (succ c) | a+b=sn = exFalso (cannotAddAndEnlarge' a+b+c<sn)
|
||||
|
||||
help26 : {a b c n : ℕ} → (a +N b ≡ n) → (a +N (b +N c) ≡ n) → 0 ≡ c
|
||||
help26 {a} {b} {c} {n} a+b=n a+b+c=n rewrite Semiring.+Associative ℕSemiring a b c | a+b=n | Semiring.commutative ℕSemiring n c = equalityCommutative (canSubtractFromEqualityRight a+b+c=n)
|
||||
|
||||
help25 : {a b c n : ℕ} → (a +N b ≡ n) → (b +N c ≡ n) → (a +N 0 ≡ c)
|
||||
help25 {a} {b} {c} {n} a+b=n b+c=n rewrite Semiring.commutative ℕSemiring a 0 | Semiring.commutative ℕSemiring a b | equalityCommutative a+b=n = equalityCommutative (canSubtractFromEqualityLeft b+c=n)
|
||||
|
||||
help24 : {a n : ℕ} → (a <N n) → (n <N a +N 0) → False
|
||||
help24 {a} {n} a<n n<a+0 rewrite Semiring.commutative ℕSemiring a 0 = TotalOrder.irreflexive ℕTotalOrder (<Transitive a<n n<a+0)
|
||||
|
||||
help23 : {a n : ℕ} → (a <N n) → (a +N 0 ≡ n) → False
|
||||
help23 {a} {n} a<n a+0=n rewrite Semiring.commutative ℕSemiring a 0 | a+0=n = TotalOrder.irreflexive ℕTotalOrder a<n
|
||||
|
||||
help22 : {a b c n : ℕ} → (a +N b ≡ n) → (c ≡ n) → (n<b+c : n <N b +N c) → (n <N a +N subtractionNResult.result (-N (inl n<b+c))) → False
|
||||
help22 {a} {b} {c} {.c} a+b=n refl n<b+c pr = TotalOrder.irreflexive ℕTotalOrder (identityOfIndiscernablesRight _<N_ pr4 a+b=n)
|
||||
where
|
||||
pr1 : c +N c <N a +N (subtractionNResult.result (-N (inl n<b+c)) +N c)
|
||||
pr1 = identityOfIndiscernablesRight _<N_ (additionPreservesInequality c pr) (equalityCommutative (Semiring.+Associative ℕSemiring a _ c))
|
||||
pr2 : c +N c <N a +N (b +N c)
|
||||
pr2 = identityOfIndiscernablesRight _<N_ pr1 (applyEquality (a +N_) (addMinus (inl n<b+c)))
|
||||
pr3 : c +N c <N (a +N b) +N c
|
||||
pr3 = identityOfIndiscernablesRight _<N_ pr2 (Semiring.+Associative ℕSemiring a b c)
|
||||
pr4 : c <N a +N b
|
||||
pr4 = subtractionPreservesInequality c pr3
|
||||
|
||||
help21 : {a b c n : ℕ} → (a +N b ≡ n) → (b +N c ≡ n) → (c ≡ n) → (a <N n) → False
|
||||
help21 {a} {b} {c} {.c} a+b=n b+c=n refl a<n = TotalOrder.irreflexive ℕTotalOrder (identityOfIndiscernablesLeft _<N_ a<n a=c)
|
||||
where
|
||||
b=0 : b ≡ 0
|
||||
a=c : a ≡ c
|
||||
a=c = identityOfIndiscernablesLeft _≡_ a+b=n lemm
|
||||
where
|
||||
lemm : a +N b ≡ a
|
||||
lemm rewrite b=0 | Semiring.commutative ℕSemiring a 0 = refl
|
||||
b=0' : (b c : ℕ) → (b +N c ≡ c) → b ≡ 0
|
||||
b=0' zero c b+c=c = refl
|
||||
b=0' (succ b) zero b+c=c = exFalso (naughtE (equalityCommutative b+c=c))
|
||||
b=0' (succ b) (succ c) b+c=c = b=0' (succ b) c bl2
|
||||
where
|
||||
bl2 : succ b +N c ≡ c
|
||||
bl2 rewrite Semiring.commutative ℕSemiring b c | Semiring.commutative ℕSemiring (succ c) b = succInjective b+c=c
|
||||
b=0 = b=0' b c b+c=n
|
||||
|
||||
help20 : {a b c n : ℕ} → (c ≡ n) → (a +N b ≡ n) → (n<b+c : n <N b +N c) → (a +N subtractionNResult.result (-N (inl n<b+c)) <N n) → False
|
||||
help20 {a} {b} {c} {n} c=n a+b=n n<b+c pr = TotalOrder.irreflexive ℕTotalOrder (identityOfIndiscernablesLeft _<N_ pr5 c=n)
|
||||
where
|
||||
pr1 : a +N (subtractionNResult.result (-N (inl n<b+c)) +N n) <N n +N n
|
||||
pr1 = identityOfIndiscernablesLeft _<N_ (additionPreservesInequality n pr) (equalityCommutative (Semiring.+Associative ℕSemiring a _ n))
|
||||
pr2 : a +N (b +N c) <N n +N n
|
||||
pr2 = identityOfIndiscernablesLeft _<N_ pr1 (applyEquality (a +N_) (addMinus (inl n<b+c)))
|
||||
pr3 : (a +N b) +N c <N n +N n
|
||||
pr3 rewrite equalityCommutative (Semiring.+Associative ℕSemiring a b c) = pr2
|
||||
pr4 : c +N n <N n +N n
|
||||
pr4 = identityOfIndiscernablesLeft _<N_ pr3 (transitivity (applyEquality (_+N c) a+b=n) (Semiring.commutative ℕSemiring n c))
|
||||
pr5 : c <N n
|
||||
pr5 = subtractionPreservesInequality n pr4
|
||||
|
||||
help19 : {a b c n : ℕ} → (b+c<n : b +N c <N n) → (n<a+b : n <N a +N b) → (a <N n) → (subtractionNResult.result (-N (inl n<a+b)) +N c ≡ n) → False
|
||||
help19 {a} {b} {c} {n} b+c<n n<a+b a<n pr = TotalOrder.irreflexive ℕTotalOrder (identityOfIndiscernablesLeft _<N_ r q')
|
||||
where
|
||||
p : (n +N subtractionNResult.result (-N (inl n<a+b))) +N c ≡ n +N n
|
||||
p = identityOfIndiscernablesLeft _≡_ (applyEquality (n +N_ ) pr) (Semiring.+Associative ℕSemiring n _ c)
|
||||
q : (a +N b) +N c ≡ n +N n
|
||||
q = identityOfIndiscernablesLeft _≡_ p (applyEquality (_+N c) (addMinus' (inl n<a+b)))
|
||||
q' : a +N (b +N c) ≡ n +N n
|
||||
q' = identityOfIndiscernablesLeft _≡_ q (equalityCommutative (Semiring.+Associative ℕSemiring a b c))
|
||||
r : a +N (b +N c) <N n +N n
|
||||
r = addStrongInequalities a<n b+c<n
|
||||
|
||||
help18 : {a b c n : ℕ} → (b+c<n : b +N c <N n) → (n<a+b : n <N a +N b) → (a <N n) → (n <N subtractionNResult.result (-N (inl n<a+b)) +N c) → False
|
||||
help18 {a} {b} {c} {n} b+c<n n<a+b a<n pr = TotalOrder.irreflexive ℕTotalOrder (<Transitive p4 a<n)
|
||||
where
|
||||
p : n +N n <N (n +N subtractionNResult.result (-N (inl n<a+b))) +N c
|
||||
p = identityOfIndiscernablesRight _<N_ (additionPreservesInequalityOnLeft n pr) (Semiring.+Associative ℕSemiring n _ c)
|
||||
p' : n +N n <N (a +N b) +N c
|
||||
p' = identityOfIndiscernablesRight _<N_ p (applyEquality (_+N c) (addMinus' (inl n<a+b)))
|
||||
p2 : n +N n <N a +N (b +N c)
|
||||
p2 = identityOfIndiscernablesRight _<N_ p' (equalityCommutative (Semiring.+Associative ℕSemiring a b c))
|
||||
p3 : n +N n <N a +N n
|
||||
p3 = <Transitive p2 (additionPreservesInequalityOnLeft a b+c<n)
|
||||
p4 : n <N a
|
||||
p4 = subtractionPreservesInequality n p3
|
||||
|
||||
help17 : {a b c n : ℕ} → (n<b+c : n <N b +N c) → (n<a+b : n <N a +N b) → (a +N subtractionNResult.result (-N (inl n<b+c)) <N n) → (subtractionNResult.result (-N (inl n<a+b)) +N c) ≡ n → False
|
||||
help17 {a} {b} {c} {n} n<b+c n<a+b p1 p2 = TotalOrder.irreflexive ℕTotalOrder (identityOfIndiscernablesLeft _<N_ pr1'' pr3)
|
||||
where
|
||||
pr1' : a +N (subtractionNResult.result (-N (inl n<b+c)) +N n) <N n +N n
|
||||
pr1' = identityOfIndiscernablesLeft _<N_ (additionPreservesInequality n p1) (equalityCommutative (Semiring.+Associative ℕSemiring a _ n))
|
||||
pr1'' : a +N (b +N c) <N n +N n
|
||||
pr1'' = identityOfIndiscernablesLeft _<N_ pr1' (applyEquality (a +N_) (addMinus (inl n<b+c)))
|
||||
pr2' : (n +N subtractionNResult.result (-N (inl n<a+b))) +N c ≡ n +N n
|
||||
pr2' = identityOfIndiscernablesLeft _≡_ (applyEquality (n +N_) p2) (Semiring.+Associative ℕSemiring n _ c)
|
||||
pr2'' : (a +N b) +N c ≡ n +N n
|
||||
pr2'' = identityOfIndiscernablesLeft _≡_ pr2' (applyEquality (_+N c) (addMinus' (inl n<a+b)))
|
||||
pr3 : a +N (b +N c) ≡ n +N n
|
||||
pr3 = identityOfIndiscernablesLeft _≡_ pr2'' (equalityCommutative (Semiring.+Associative ℕSemiring a b c))
|
||||
|
||||
help16 : {a b c n : ℕ} → (n<b+c : n <N b +N c) → (n<a+b : n <N a +N b) → (a +N subtractionNResult.result (-N (inl n<b+c))) <N n → (pr : n <N subtractionNResult.result (-N (inl n<a+b)) +N c) → a +N subtractionNResult.result (-N (inl n<b+c)) ≡ subtractionNResult.result (-N (inl pr))
|
||||
help16 {a} {b} {c} {n} n<b+c n<a+b pr1 pr2 = exFalso (TotalOrder.irreflexive ℕTotalOrder (<Transitive pr3 pr1''))
|
||||
where
|
||||
pr1' : a +N (subtractionNResult.result (-N (inl n<b+c)) +N n) <N n +N n
|
||||
pr1' = identityOfIndiscernablesLeft _<N_ (additionPreservesInequality n pr1) (equalityCommutative (Semiring.+Associative ℕSemiring a _ n))
|
||||
pr1'' : a +N (b +N c) <N n +N n
|
||||
pr1'' = identityOfIndiscernablesLeft _<N_ pr1' (applyEquality (a +N_) (addMinus (inl n<b+c)))
|
||||
pr2' : n +N n <N (n +N subtractionNResult.result (-N (inl n<a+b))) +N c
|
||||
pr2' = identityOfIndiscernablesRight _<N_ (additionPreservesInequalityOnLeft n pr2) (Semiring.+Associative ℕSemiring n _ c)
|
||||
pr2'' : n +N n <N (a +N b) +N c
|
||||
pr2'' = identityOfIndiscernablesRight _<N_ pr2' (applyEquality (_+N c) (addMinus' (inl n<a+b)))
|
||||
pr3 : n +N n <N a +N (b +N c)
|
||||
pr3 = identityOfIndiscernablesRight _<N_ pr2'' (equalityCommutative (Semiring.+Associative ℕSemiring a b c))
|
||||
|
||||
help15 : {a b c n : ℕ} → (n<b+c : n <N b +N c) → (n<a+b : n <N a +N b) → (n <N a +N subtractionNResult.result (-N (inl n<b+c))) → (subtractionNResult.result (-N (inl n<a+b)) +N c) <N n → False
|
||||
help15 {a} {b} {c} {n} n<b+c n<a+b pr1 pr2 = TotalOrder.irreflexive ℕTotalOrder (<Transitive p2'' p1')
|
||||
where
|
||||
p1 : (n +N subtractionNResult.result (-N (inl n<a+b))) +N c <N n +N n
|
||||
p1 = identityOfIndiscernablesLeft _<N_ (additionPreservesInequalityOnLeft n pr2) (Semiring.+Associative ℕSemiring n _ c)
|
||||
p1' : (a +N b) +N c <N n +N n
|
||||
p1' = identityOfIndiscernablesLeft _<N_ p1 (applyEquality (_+N c) (addMinus' (inl n<a+b)))
|
||||
p2 : n +N n <N a +N (subtractionNResult.result (-N (inl n<b+c)) +N n)
|
||||
p2 = identityOfIndiscernablesRight _<N_ (additionPreservesInequality n pr1) (equalityCommutative (Semiring.+Associative ℕSemiring a _ n))
|
||||
p2' : n +N n <N a +N (b +N c)
|
||||
p2' = identityOfIndiscernablesRight _<N_ p2 (applyEquality (a +N_) (addMinus (inl n<b+c)))
|
||||
p2'' : n +N n <N (a +N b) +N c
|
||||
p2'' = identityOfIndiscernablesRight _<N_ p2' (Semiring.+Associative ℕSemiring a b c)
|
||||
|
||||
help14 : {a b c n : ℕ} → (n<b+c : n <N b +N c) → (n<a+b : n <N a +N b) → (pr1 : n <N a +N subtractionNResult.result (-N (inl n<b+c))) → (pr2 : n <N subtractionNResult.result (-N (inl n<a+b)) +N c) → subtractionNResult.result (-N (inl pr1)) ≡ subtractionNResult.result (-N (inl pr2))
|
||||
help14 {a} {b} {c} {n} n<b+c n<a+b pr1 pr2 = equivalentSubtraction _ _ _ _ pr1 pr2 (transitivity (Semiring.+Associative ℕSemiring n _ c) (transitivity (applyEquality (_+N c) (addMinus' (inl n<a+b))) (transitivity (equalityCommutative (Semiring.+Associative ℕSemiring a b c)) (equalityCommutative p2))))
|
||||
where
|
||||
p1 : (a +N subtractionNResult.result (-N (inl n<b+c))) +N n ≡ a +N (subtractionNResult.result (-N (inl n<b+c)) +N n)
|
||||
p1 = equalityCommutative (Semiring.+Associative ℕSemiring a _ n)
|
||||
p2 : (a +N subtractionNResult.result (-N (inl n<b+c))) +N n ≡ a +N (b +N c)
|
||||
p2 = identityOfIndiscernablesRight _≡_ p1 (applyEquality (a +N_) (addMinus (inl n<b+c)))
|
||||
|
||||
help13 : {a b c n : ℕ} → (n<b+c : n <N b +N c) → (n<a+b : n <N a +N b) → (n <N a +N subtractionNResult.result (-N (inl n<b+c))) → (subtractionNResult.result (-N (inl n<a+b)) +N c ≡ n) → False
|
||||
help13 {a} {b} {c} {n} n<b+c n<a+b pr1 pr2 = TotalOrder.irreflexive ℕTotalOrder (identityOfIndiscernablesRight _<N_ lemm1' lemm3)
|
||||
where
|
||||
lemm1 : n +N n <N a +N (subtractionNResult.result (-N (inl n<b+c)) +N n)
|
||||
lemm1 = identityOfIndiscernablesRight _<N_ (additionPreservesInequality n pr1) (equalityCommutative (Semiring.+Associative ℕSemiring a _ n))
|
||||
lemm1' : n +N n <N a +N (b +N c)
|
||||
lemm1' = identityOfIndiscernablesRight _<N_ lemm1 (applyEquality (a +N_) (addMinus (inl n<b+c)))
|
||||
lemm2 : (n +N subtractionNResult.result (-N (inl n<a+b))) +N c ≡ n +N n
|
||||
lemm2 = identityOfIndiscernablesLeft _≡_ (applyEquality (n +N_) pr2) (Semiring.+Associative ℕSemiring n _ c)
|
||||
lemm2' : (a +N b) +N c ≡ n +N n
|
||||
lemm2' = identityOfIndiscernablesLeft _≡_ lemm2 (applyEquality (_+N c) (addMinus' (inl n<a+b)))
|
||||
lemm3 : a +N (b +N c) ≡ n +N n
|
||||
lemm3 rewrite Semiring.+Associative ℕSemiring a b c = lemm2'
|
||||
|
||||
help12 : {a b c n : ℕ} → (n<b+c : n <N b +N c) → (n<a+b : n <N a +N b) → (a +N subtractionNResult.result (-N (inl n<b+c))) ≡ n → subtractionNResult.result (-N (inl n<a+b)) +N c <N n → False
|
||||
help12 {a} {b} {c} {n} n<b+c n<a+b pr1 pr2 = TotalOrder.irreflexive ℕTotalOrder lemm4
|
||||
where
|
||||
pr : {a b c : ℕ} → a +N (b +N c) ≡ b +N (a +N c)
|
||||
pr {a} {b} {c} rewrite Semiring.+Associative ℕSemiring a b c | Semiring.commutative ℕSemiring a b | equalityCommutative (Semiring.+Associative ℕSemiring b a c) = refl
|
||||
lemm1 : (n +N subtractionNResult.result (-N (inl n<a+b))) +N c <N n +N n
|
||||
lemm1 = identityOfIndiscernablesLeft _<N_ (additionPreservesInequalityOnLeft n pr2) (Semiring.+Associative ℕSemiring n _ c)
|
||||
lemm2 : (a +N b) +N c <N n +N n
|
||||
lemm2 = identityOfIndiscernablesLeft _<N_ lemm1 (applyEquality (_+N c) (addMinus' (inl n<a+b)))
|
||||
lemm1' : a +N (subtractionNResult.result (-N (inl n<b+c)) +N n) ≡ n +N n
|
||||
lemm1' = identityOfIndiscernablesLeft _≡_ (applyEquality (_+N n) pr1) (equalityCommutative (Semiring.+Associative ℕSemiring a _ n))
|
||||
lemm2' : a +N (b +N c) ≡ n +N n
|
||||
lemm2' = identityOfIndiscernablesLeft _≡_ lemm1' (applyEquality (a +N_) (addMinus (inl n<b+c)))
|
||||
lemm3 : (a +N b) +N c ≡ n +N n
|
||||
lemm3 rewrite equalityCommutative (Semiring.+Associative ℕSemiring a b c) = lemm2'
|
||||
lemm4 : (a +N b) +N c <N (a +N b) +N c
|
||||
lemm4 = identityOfIndiscernablesRight _<N_ lemm2 (equalityCommutative lemm3)
|
||||
|
||||
help11 : {a b c n : ℕ} → (a <N n) → (b +N c ≡ n) → (n<a+b : n <N a +N b) → (n <N subtractionNResult.result (-N (inl n<a+b)) +N c) → False
|
||||
help11 {a} {b} {c} {n} a<n b+c=n n<a+b pr1 = TotalOrder.irreflexive ℕTotalOrder (<Transitive a<n lemm5)
|
||||
where
|
||||
pr : {a b c : ℕ} → a +N (b +N c) ≡ b +N (a +N c)
|
||||
pr {a} {b} {c} rewrite Semiring.+Associative ℕSemiring a b c | Semiring.commutative ℕSemiring a b | equalityCommutative (Semiring.+Associative ℕSemiring b a c) = refl
|
||||
lemm : n +N n <N (n +N subtractionNResult.result (-N (inl n<a+b))) +N c
|
||||
lemm = identityOfIndiscernablesRight _<N_ (additionPreservesInequalityOnLeft n pr1) (Semiring.+Associative ℕSemiring n _ c)
|
||||
lemm2 : n +N n <N (a +N b) +N c
|
||||
lemm2 = identityOfIndiscernablesRight _<N_ lemm (applyEquality (_+N c) (addMinus' (inl n<a+b)))
|
||||
lemm3 : n +N n <N a +N (b +N c)
|
||||
lemm3 = identityOfIndiscernablesRight _<N_ lemm2 (equalityCommutative (Semiring.+Associative ℕSemiring a b c))
|
||||
lemm4 : n +N n <N a +N n
|
||||
lemm4 = identityOfIndiscernablesRight _<N_ lemm3 (applyEquality (a +N_) b+c=n)
|
||||
lemm5 : n <N a
|
||||
lemm5 = subtractionPreservesInequality n lemm4
|
||||
|
||||
help10 : {a b c n : ℕ} → (n<b+c : n <N b +N c) → (n<a+b : n <N a +N b) → (a +N subtractionNResult.result (-N (inl n<b+c)) ≡ n) → (n <N subtractionNResult.result (-N (inl n<a+b)) +N c) → False
|
||||
help10 {a} {b} {c} {n} n<b+c n<a+b pr1 pr2 = TotalOrder.irreflexive ℕTotalOrder lemm6
|
||||
where
|
||||
pr : {a b c : ℕ} → a +N (b +N c) ≡ b +N (a +N c)
|
||||
pr {a} {b} {c} rewrite Semiring.+Associative ℕSemiring a b c | Semiring.commutative ℕSemiring a b | equalityCommutative (Semiring.+Associative ℕSemiring b a c) = refl
|
||||
lemm : a +N (n +N subtractionNResult.result (-N (inl n<b+c))) ≡ n +N n
|
||||
lemm = identityOfIndiscernablesLeft _≡_ (applyEquality (n +N_) pr1) (pr {n} {a})
|
||||
lemm2 : a +N (b +N c) ≡ n +N n
|
||||
lemm2 = identityOfIndiscernablesLeft _≡_ lemm (applyEquality (a +N_) (addMinus' (inl n<b+c)))
|
||||
lemm3 : n +N n <N (n +N subtractionNResult.result (-N (inl n<a+b))) +N c
|
||||
lemm3 = identityOfIndiscernablesRight _<N_ (additionPreservesInequalityOnLeft n pr2) (Semiring.+Associative ℕSemiring n _ c)
|
||||
lemm4 : n +N n <N (a +N b) +N c
|
||||
lemm4 = identityOfIndiscernablesRight _<N_ lemm3 (applyEquality (_+N c) (addMinus' (inl n<a+b)))
|
||||
lemm5 : n +N n <N a +N (b +N c)
|
||||
lemm5 = identityOfIndiscernablesRight _<N_ lemm4 (equalityCommutative (Semiring.+Associative ℕSemiring a b c))
|
||||
lemm6 : a +N (b +N c) <N a +N (b +N c)
|
||||
lemm6 = identityOfIndiscernablesLeft _<N_ lemm5 (equalityCommutative lemm2)
|
||||
|
||||
help9 : {a n : ℕ} → (a +N 0 ≡ n) → (a <N n) → False
|
||||
help9 {a} {n} n=a+0 a<n rewrite Semiring.commutative ℕSemiring a 0 | n=a+0 = TotalOrder.irreflexive ℕTotalOrder a<n
|
||||
|
||||
help8 : {a n : ℕ} → (n <N a +N 0) → (a <N n) → False
|
||||
help8 {a} {n} n<a+0 a<n rewrite Semiring.commutative ℕSemiring a 0 = TotalOrder.irreflexive ℕTotalOrder (<Transitive a<n n<a+0)
|
||||
|
||||
help6 : {a b c n : ℕ} → (b +N c ≡ n) → (n<a+b : n <N a +N b) → (a +N 0 ≡ subtractionNResult.result (-N (inl n<a+b)) +N c)
|
||||
help6 {a} {b} {c} {n} b+c=n n<a+b rewrite Semiring.commutative ℕSemiring a 0 = canSubtractFromEqualityLeft {n} lem'
|
||||
where
|
||||
lem : n +N a ≡ (n +N subtractionNResult.result (-N (inl n<a+b))) +N c
|
||||
lem rewrite addMinus' (inl n<a+b) | equalityCommutative (Semiring.+Associative ℕSemiring a b c) | b+c=n = Semiring.commutative ℕSemiring n a
|
||||
lem' : n +N a ≡ n +N (subtractionNResult.result (-N (inl n<a+b)) +N c)
|
||||
lem' = identityOfIndiscernablesRight _≡_ lem (equalityCommutative (Semiring.+Associative ℕSemiring n _ c))
|
||||
|
||||
help5 : {a b c n : ℕ} → (n<b+c : n <N b +N c) → (n<a+b : n <N a +N b) → a +N subtractionNResult.result (-N (inl n<b+c)) ≡ subtractionNResult.result (-N (inl n<a+b)) +N c
|
||||
help5 {a} {b} {c} {n} n<b+c n<a+b = canSubtractFromEqualityLeft {n} lemma''
|
||||
where
|
||||
lemma : a +N (n +N subtractionNResult.result (-N (inl n<b+c))) ≡ (n +N subtractionNResult.result (-N (inl n<a+b))) +N c
|
||||
lemma rewrite addMinus' (inl n<b+c) | addMinus' (inl n<a+b) = Semiring.+Associative ℕSemiring a b c
|
||||
lemma' : a +N (n +N subtractionNResult.result (-N (inl n<b+c))) ≡ n +N (subtractionNResult.result (-N (inl n<a+b)) +N c)
|
||||
lemma' = identityOfIndiscernablesRight _≡_ lemma (equalityCommutative (Semiring.+Associative ℕSemiring n _ c))
|
||||
lemma'' : n +N (a +N subtractionNResult.result (-N (inl n<b+c))) ≡ n +N (subtractionNResult.result (-N (inl n<a+b)) +N c)
|
||||
lemma'' = identityOfIndiscernablesLeft _≡_ lemma' (pr {a} {n} {subtractionNResult.result (-N (inl n<b+c))})
|
||||
where
|
||||
pr : {a b c : ℕ} → a +N (b +N c) ≡ b +N (a +N c)
|
||||
pr {a} {b} {c} rewrite Semiring.+Associative ℕSemiring a b c | Semiring.commutative ℕSemiring a b | equalityCommutative (Semiring.+Associative ℕSemiring b a c) = refl
|
||||
|
||||
help4 : {a b c n : ℕ} → (n<a+'b+c : n <N a +N (b +N c)) → (n<a+b : n <N a +N b) → (subtractionNResult.result (-N (inl n<a+'b+c)) ≡ subtractionNResult.result (-N (inl n<a+b)) +N c)
|
||||
help4 {a} {b} {c} {n} n<a+'b+c n<a+b = canSubtractFromEqualityLeft lemma'
|
||||
where
|
||||
lemma : (n +N subtractionNResult.result (-N (inl n<a+'b+c))) ≡ (n +N subtractionNResult.result (-N (inl n<a+b))) +N c
|
||||
lemma rewrite addMinus' (inl n<a+'b+c) | addMinus' (inl n<a+b) = Semiring.+Associative ℕSemiring a b c
|
||||
lemma' : n +N subtractionNResult.result (-N (inl n<a+'b+c)) ≡ n +N (subtractionNResult.result (-N (inl n<a+b)) +N c)
|
||||
lemma' = identityOfIndiscernablesRight _≡_ lemma (equalityCommutative (Semiring.+Associative ℕSemiring n (subtractionNResult.result (-N (inl n<a+b))) c))
|
||||
|
||||
help3 : {a b c n : ℕ} → (a <N n) → (b <N n) → (c <N n) → (a +N b <N n) → (pr : n <N b +N c) → a +N subtractionNResult.result (-N (inl pr)) ≡ n → False
|
||||
help3 {a} {b} {c} {n} a<n b<n c<n a+b<n n<b+c pr = TotalOrder.irreflexive ℕTotalOrder (<Transitive (inter4 inter3) c<n)
|
||||
where
|
||||
inter : a +N (n +N subtractionNResult.result (-N (inl n<b+c))) ≡ n +N n
|
||||
inter = identityOfIndiscernablesLeft _≡_ (applyEquality (n +N_) pr) (lemma n a (subtractionNResult.result (-N (inl n<b+c))))
|
||||
where
|
||||
lemma : (a b c : ℕ) → a +N (b +N c) ≡ b +N (a +N c)
|
||||
lemma a b c rewrite Semiring.+Associative ℕSemiring a b c | Semiring.+Associative ℕSemiring b a c = applyEquality (_+N c) (Semiring.commutative ℕSemiring a b)
|
||||
inter2 : n +N n ≡ a +N (b +N c)
|
||||
inter2 = equalityCommutative (identityOfIndiscernablesLeft _≡_ inter (applyEquality (a +N_) (addMinus' (inl n<b+c))))
|
||||
inter3 : n +N n <N n +N c
|
||||
inter3 rewrite inter2 | Semiring.+Associative ℕSemiring a b c = additionPreservesInequality c a+b<n
|
||||
inter4 : (n +N n <N n +N c) → n <N c
|
||||
inter4 pr rewrite Semiring.commutative ℕSemiring n c = subtractionPreservesInequality n pr
|
||||
|
||||
help2 : {a b c n : ℕ} → (sn<b+c : succ n <N b +N c) → (sn<a+b+c : succ n <N (a +N b) +N c) → a +N subtractionNResult.result (-N (inl sn<b+c)) ≡ subtractionNResult.result (-N (inl sn<a+b+c))
|
||||
help2 {a} {b} {c} {n} sn<b+c sn<a+b+c = res inter
|
||||
where
|
||||
inter : a +N (subtractionNResult.result (-N (inl sn<b+c)) +N succ n) ≡ subtractionNResult.result (-N (inl sn<a+b+c)) +N succ n
|
||||
inter rewrite addMinus (inl sn<b+c) | addMinus (inl sn<a+b+c) = Semiring.+Associative ℕSemiring a b c
|
||||
res : (a +N (subtractionNResult.result (-N (inl sn<b+c)) +N succ n) ≡ subtractionNResult.result (-N (inl sn<a+b+c)) +N succ n) → a +N subtractionNResult.result (-N (inl sn<b+c)) ≡ subtractionNResult.result (-N (inl sn<a+b+c))
|
||||
res pr = canSubtractFromEqualityRight {_} {succ n} (identityOfIndiscernablesLeft _≡_ pr (Semiring.+Associative ℕSemiring a (subtractionNResult.result (-N (inl sn<b+c))) (succ n)))
|
||||
|
||||
help1 : {a b c n : ℕ} → (sn<b+c : succ n <N b +N c) → (pr1 : succ n <N a +N subtractionNResult.result (-N (inl sn<b+c))) → (a +N b <N succ n) → (a <N succ n) → (b <N succ n) → (c <N succ n) → False
|
||||
help1 {a} {b} {c} {n} sn<b+c pr1 a+b<sn a<sn b<sn c<sn with -N (inl sn<b+c)
|
||||
help1 {a} {b} {c} {n} sn<b+c pr1 a+b<sn a<sn b<sn c<sn | record { result = b+c-sn ; pr = Prb+c-sn } = ans
|
||||
where
|
||||
b+c-nNonzero : b+c-sn ≡ 0 → False
|
||||
b+c-nNonzero pr rewrite (equalityCommutative Prb+c-sn) | pr | Semiring.commutative ℕSemiring n 0 = TotalOrder.irreflexive ℕTotalOrder sn<b+c
|
||||
2sn<a+b+c' : succ n +N succ n <N succ n +N (a +N b+c-sn)
|
||||
2sn<a+b+c' = additionPreservesInequalityOnLeft (succ n) pr1
|
||||
2sn<a+b+c'' : succ n +N succ n <N a +N (succ n +N b+c-sn)
|
||||
2sn<a+b+c'' rewrite Semiring.+Associative ℕSemiring a (succ n) b+c-sn | Semiring.commutative ℕSemiring a (succ n) | equalityCommutative (Semiring.+Associative ℕSemiring (succ n) a b+c-sn) = 2sn<a+b+c'
|
||||
eep : succ n +N succ n <N a +N (b +N c)
|
||||
eep rewrite equalityCommutative Prb+c-sn = 2sn<a+b+c''
|
||||
eep2 : a +N (b +N c) <N succ n +N c
|
||||
eep2 rewrite Semiring.+Associative ℕSemiring a b c = additionPreservesInequality c a+b<sn
|
||||
eep2' : a +N (b +N c) <N succ n +N succ n
|
||||
eep2' = <Transitive eep2 (additionPreservesInequalityOnLeft (succ n) c<sn)
|
||||
ans : False
|
||||
ans = TotalOrder.irreflexive ℕTotalOrder (<Transitive eep eep2')
|
||||
|
||||
plusZnAssociative : {n : ℕ} → {pr : 0 <N n} → (a b c : ℤn n pr) → a +n (b +n c) ≡ ((a +n b) +n c)
|
||||
plusZnAssociative {zero} {()}
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess} record { x = c ; xLess = cLess } with totality (a +N b) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) with totality ((a +N b) +N c) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) with totality (b +N c) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inl b+c<sn) with totality (a +N (b +N c)) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inl b+c<sn) | inl (inl a+'b+c<sn) = equalityZn _ _ (Semiring.+Associative ℕSemiring a b c)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inl b+c<sn) | inl (inr sn<a+'b+c) = exFalso (false {succ n} a+b+c<sn sn<a+'b+c)
|
||||
where
|
||||
false : {x : ℕ} → (a +N b) +N c <N succ n → succ n <N a +N (b +N c) → False
|
||||
false pr1 pr2 rewrite equalityCommutative (Semiring.+Associative ℕSemiring a b c) = TotalOrder.irreflexive ℕTotalOrder (<Transitive pr1 pr2)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inl b+c<sn) | inr sn=a+b+c = exFalso (false a+b+c<sn sn=a+b+c)
|
||||
where
|
||||
false : {x : ℕ} → (a +N b) +N c <N x → (a +N (b +N c)) ≡ x → False
|
||||
false p1 p2 rewrite equalityCommutative (Semiring.+Associative ℕSemiring a b c) | p2 = TotalOrder.irreflexive ℕTotalOrder p1
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inr sn<b+c) with totality (a +N (b +N c)) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inr sn<b+c) | inl (inl _) with totality (a +N subtractionNResult.result (-N (inl sn<b+c))) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inr sn<b+c) | inl (inl _) | inl (inl x) with -N (inl sn<b+c)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inr sn<b+c) | inl (inl a+'b+c<sn) | inl (inl x) | record { result = result ; pr = pr } = exFalso (false a+'b+c<sn pr)
|
||||
where
|
||||
false : a +N (b +N c) <N succ n → succ n +N result ≡ b +N c → False
|
||||
false pr1 pr2 rewrite equalityCommutative pr2 | Semiring.+Associative ℕSemiring a (succ n) result | Semiring.commutative ℕSemiring a (succ n) | equalityCommutative (Semiring.+Associative ℕSemiring (succ n) a result) = cannotAddAndEnlarge' pr1
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inr sn<b+c) | inl (inl _) | inl (inr x) with -N (inl sn<b+c)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inr sn<b+c) | inl (inl a+'b+c<sn) | inl (inr x) | record { result = result ; pr = pr } = exFalso false
|
||||
where
|
||||
lemma : a +N (succ n +N result) ≡ a +N (b +N c)
|
||||
lemma' : a +N (succ n +N result) <N succ n
|
||||
lemma'' : succ n +N (a +N result) <N succ n
|
||||
lemma'' = identityOfIndiscernablesLeft _<N_ lemma' (transitivity (Semiring.+Associative ℕSemiring a (succ n) result) (transitivity (applyEquality (λ t → t +N result) (Semiring.commutative ℕSemiring a (succ n))) (equalityCommutative (Semiring.+Associative ℕSemiring (succ n) a result))))
|
||||
lemma = applyEquality (λ i → a +N i) pr
|
||||
lemma' rewrite lemma = a+'b+c<sn
|
||||
false : False
|
||||
false = cannotAddAndEnlarge' lemma''
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inr sn<b+c) | inl (inl _) | inr x with -N (inl sn<b+c)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inr sn<b+c) | inl (inl a+'b+c<sn) | inr x | record { result = result ; pr = pr } = exFalso false
|
||||
where
|
||||
lemma : a +N (succ n +N result) <N succ n
|
||||
lemma rewrite pr = a+'b+c<sn
|
||||
lemma' : succ n +N (a +N result) <N succ n
|
||||
lemma' = identityOfIndiscernablesLeft _<N_ lemma (transitivity (Semiring.+Associative ℕSemiring a (succ n) result) (transitivity (applyEquality (λ t → t +N result) (Semiring.commutative ℕSemiring a (succ n))) (equalityCommutative (Semiring.+Associative ℕSemiring (succ n) a result))))
|
||||
false : False
|
||||
false = cannotAddAndEnlarge' lemma'
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inr sn<b+c) | inl (inr x) = equalityZn _ _ (exFalso (false {succ n} a+b+c<sn x))
|
||||
where
|
||||
false : {x : ℕ} → (a +N b) +N c <N succ n → succ n <N a +N (b +N c) → False
|
||||
false pr1 pr2 rewrite equalityCommutative (Semiring.+Associative ℕSemiring a b c) = TotalOrder.irreflexive ℕTotalOrder (<Transitive pr1 pr2)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inr sn<b+c) | inr x = exFalso (false a+b+c<sn x)
|
||||
where
|
||||
false : (a +N b) +N c <N succ n → a +N (b +N c) ≡ succ n → False
|
||||
false pr1 pr2 rewrite equalityCommutative pr2 | equalityCommutative (Semiring.+Associative ℕSemiring a b c) = TotalOrder.irreflexive ℕTotalOrder pr1
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inr sn=b+c = exFalso (false a+b+c<sn sn=b+c)
|
||||
where
|
||||
false : (a +N b) +N c <N succ n → b +N c ≡ succ n → False
|
||||
false pr1 pr2 rewrite equalityCommutative (Semiring.+Associative ℕSemiring a b c) | pr2 | Semiring.commutative ℕSemiring a (succ n) = cannotAddAndEnlarge' a+b+c<sn
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) with totality (b +N c) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inl (inl b+c<sn) with totality (a +N (b +N c)) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inl (inl b+c<sn) | inl (inl a+'b+c<sn) = exFalso (false sn<a+b+c a+'b+c<sn)
|
||||
where
|
||||
false : (succ n <N (a +N b) +N c) → a +N (b +N c) <N succ n → False
|
||||
false pr1 pr2 rewrite equalityCommutative (Semiring.+Associative ℕSemiring a b c) = TotalOrder.irreflexive ℕTotalOrder (<Transitive pr1 pr2)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inl (inl b+c<sn) | inl (inr sn<a+'b+c) = equalityZn _ _ ans
|
||||
where
|
||||
lemma : succ n +N ((a +N b) +N c) ≡ (a +N (b +N c)) +N succ n
|
||||
lemma rewrite equalityCommutative (Semiring.+Associative ℕSemiring a b c) = Semiring.commutative ℕSemiring (succ n) _
|
||||
ans : subtractionNResult.result (-N (inl sn<a+'b+c)) ≡ subtractionNResult.result (-N (inl sn<a+b+c))
|
||||
ans = equivalentSubtraction _ _ _ _ sn<a+'b+c sn<a+b+c lemma
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inl (inl b+c<sn) | inr x = exFalso (false sn<a+b+c x)
|
||||
where
|
||||
false : succ n <N (a +N b) +N c → (a +N (b +N c)) ≡ succ n → False
|
||||
false pr1 pr2 rewrite equalityCommutative (Semiring.+Associative ℕSemiring a b c) | pr2 = TotalOrder.irreflexive ℕTotalOrder sn<a+b+c
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inl (inr sn<b+c) with totality (a +N (b +N c)) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inl (inr sn<b+c) | inl (inl a+b+c<sn) = exFalso (false sn<a+b+c a+b+c<sn)
|
||||
where
|
||||
false : succ n <N (a +N b) +N c → (a +N (b +N c)) <N succ n → False
|
||||
false pr1 pr2 rewrite equalityCommutative (Semiring.+Associative ℕSemiring a b c) = TotalOrder.irreflexive ℕTotalOrder (<Transitive pr1 pr2)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inl (inr sn<b+c) | inl (inr _) with totality (a +N subtractionNResult.result (-N (inl sn<b+c))) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inl (inr sn<b+c) | inl (inr _) | inl (inl x) = equalityZn _ _ (help2 {a} {b} {c} {n} sn<b+c sn<a+b+c)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inl (inr sn<b+c) | inl (inr _) | inl (inr x) = equalityZn _ _ (exFalso (help1 sn<b+c x a+b<sn aLess bLess cLess))
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inl (inr sn<b+c) | inl (inr _) | inr x = exFalso (help3 aLess bLess cLess a+b<sn sn<b+c x)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inl (inr sn<b+c) | inr a+b+c=sn = exFalso (false sn<a+b+c a+b+c=sn)
|
||||
where
|
||||
false : (succ n <N (a +N b) +N c) → (a +N (b +N c) ≡ succ n) → False
|
||||
false pr1 pr2 rewrite equalityCommutative (Semiring.+Associative ℕSemiring a b c) | pr2 = TotalOrder.irreflexive ℕTotalOrder pr1
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inr sn=b+c with totality (a +N (b +N c)) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inr sn=b+c | inl (inl x) = exFalso (false sn<a+b+c x)
|
||||
where
|
||||
false : succ n <N (a +N b) +N c → a +N (b +N c) <N succ n → False
|
||||
false p1 p2 rewrite equalityCommutative (Semiring.+Associative ℕSemiring a b c) = TotalOrder.irreflexive ℕTotalOrder (<Transitive p1 p2)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inr sn=b+c | inl (inr _) with totality (a +N 0) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inr sn=b+c | inl (inr _) | inl (inl x) = equalityZn _ _ (ans sn=b+c)
|
||||
where
|
||||
ans : b +N c ≡ succ n → a +N 0 ≡ subtractionNResult.result (-N (inl sn<a+b+c))
|
||||
ans pr with -N (inl sn<a+b+c)
|
||||
ans b+c=sn | record { result = result ; pr = pr1 } rewrite Semiring.commutative ℕSemiring a 0 | equalityCommutative (Semiring.+Associative ℕSemiring a b c) | b+c=sn | Semiring.commutative ℕSemiring (succ n) result = equalityCommutative (canSubtractFromEqualityRight pr1)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inr sn=b+c | inl (inr _) | inl (inr x) = exFalso (false b a+b<sn x)
|
||||
where
|
||||
false : (b : ℕ) → a +N b <N succ n → succ n <N a +N 0 → False
|
||||
false zero pr1 pr2 rewrite Semiring.commutative ℕSemiring a 0 = TotalOrder.irreflexive ℕTotalOrder (<Transitive pr1 pr2)
|
||||
false (succ b) pr1 pr2 rewrite Semiring.commutative ℕSemiring a 0 = TotalOrder.irreflexive ℕTotalOrder (<Transitive pr2 (<Transitive (le b (Semiring.commutative ℕSemiring (succ b) a)) pr1))
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inr sn=b+c | inl (inr _) | inr x = exFalso (false b a+b<sn x)
|
||||
where
|
||||
false : (b : ℕ) → a +N b <N succ n → a +N 0 ≡ succ n → False
|
||||
false zero pr1 pr2 rewrite pr2 = TotalOrder.irreflexive ℕTotalOrder pr1
|
||||
false (succ b) pr1 pr2 rewrite Semiring.commutative ℕSemiring a 0 | pr2 = cannotAddAndEnlarge' pr1
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inr sn=b+c | inr x = exFalso (false sn<a+b+c x)
|
||||
where
|
||||
false : succ n <N (a +N b) +N c → a +N (b +N c) ≡ succ n → False
|
||||
false pr1 pr2 rewrite equalityCommutative (Semiring.+Associative ℕSemiring a b c) | pr2 = TotalOrder.irreflexive ℕTotalOrder pr1
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inr sn=a+b+c with totality (b +N c) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inr sn=a+b+c | inl (inl b+c<sn) with totality (a +N (b +N c)) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inr sn=a+b+c | inl (inl b+c<sn) | inl (inl a+'b+c<sn) = exFalso (false sn=a+b+c a+'b+c<sn)
|
||||
where
|
||||
false : (a +N b) +N c ≡ succ n → a +N (b +N c) <N succ n → False
|
||||
false pr1 pr2 rewrite equalityCommutative (Semiring.+Associative ℕSemiring a b c) | pr1 = TotalOrder.irreflexive ℕTotalOrder pr2
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inr sn=a+b+c | inl (inl b+c<sn) | inl (inr sn<a+'b+c) = exFalso (false sn=a+b+c sn<a+'b+c)
|
||||
where
|
||||
false : (a +N b) +N c ≡ succ n → succ n <N a +N (b +N c) → False
|
||||
false pr1 pr2 rewrite equalityCommutative (Semiring.+Associative ℕSemiring a b c) | pr1 = TotalOrder.irreflexive ℕTotalOrder pr2
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inr sn=a+b+c | inl (inl b+c<sn) | inr _ = equalityZn _ _ refl
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inr sn=a+b+c | inl (inr sn<b+c) = exFalso (false a sn=a+b+c sn<b+c)
|
||||
where
|
||||
false : (a : ℕ) → (a +N b) +N c ≡ succ n → succ n <N b +N c → False
|
||||
false zero pr1 pr2 rewrite equalityCommutative (Semiring.+Associative ℕSemiring a b c) | equalityCommutative pr1 = TotalOrder.irreflexive ℕTotalOrder pr2
|
||||
false (succ a) pr1 pr2 rewrite equalityCommutative (Semiring.+Associative ℕSemiring a b c) | equalityCommutative pr1 = TotalOrder.irreflexive ℕTotalOrder (<Transitive pr2 (le a refl))
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inr sn=a+b+c | inr b+c=sn with totality (a +N 0) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inr sn=a+b+c | inr b+c=sn | inl (inl a+0<sn) = equalityZn _ _ ans
|
||||
where
|
||||
a=0 : (a : ℕ) → (a +N b) +N c ≡ succ n → b +N c ≡ succ n → a ≡ 0
|
||||
a=0 zero pr1 pr2 rewrite equalityCommutative (Semiring.+Associative ℕSemiring a b c) | pr2 = refl
|
||||
a=0 (succ a) pr1 pr2 rewrite equalityCommutative (Semiring.+Associative ℕSemiring a b c) | pr2 = canSubtractFromEqualityRight pr1
|
||||
ans : a +N 0 ≡ 0
|
||||
ans rewrite Semiring.commutative ℕSemiring a 0 = a=0 a sn=a+b+c b+c=sn
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inr sn=a+b+c | inr b+c=sn | inl (inr sn<a+0) = exFalso (false sn<a+0 sn=a+b+c b+c=sn)
|
||||
where
|
||||
false : succ n <N a +N 0 → (a +N b) +N c ≡ succ n → b +N c ≡ succ n → False
|
||||
false pr1 pr2 pr3 rewrite equalityCommutative (Semiring.+Associative ℕSemiring a b c) | pr3 | Semiring.commutative ℕSemiring a 0 = zeroNeverGreater {succ n} (identityOfIndiscernablesRight _<N_ pr1 (a=0 a pr2))
|
||||
where
|
||||
a=0 : (a : ℕ) → (a +N succ n ≡ succ n) → a ≡ 0
|
||||
a=0 zero pr = refl
|
||||
a=0 (succ a) pr = canSubtractFromEqualityRight pr
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inr sn=a+b+c | inr b+c=sn | inr sn=a+0 = exFalso (false sn=a+b+c b+c=sn sn=a+0)
|
||||
where
|
||||
false : (a +N b) +N c ≡ succ n → b +N c ≡ succ n → a +N 0 ≡ succ n → False
|
||||
false pr1 pr2 pr3 rewrite equalityCommutative (Semiring.+Associative ℕSemiring a b c) | pr2 | equalityCommutative pr3 | Semiring.commutative ℕSemiring a 0 = naughtE (identityOfIndiscernablesLeft _≡_ pr3 (a=0 a pr1))
|
||||
where
|
||||
a=0 : (a : ℕ) → (a +N a ≡ a) → a ≡ 0
|
||||
a=0 zero pr = refl
|
||||
a=0 (succ a) pr = exFalso (naughtE {a} (equalityCommutative (canSubtractFromEqualityRight pr)))
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) with totality (b +N c) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inl (inl b+c<sn) with totality (a +N (b +N c)) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inl (inl b+c<sn) | inl (inl a+'b+c<sn) = exFalso (false sn<a+b a+'b+c<sn)
|
||||
where
|
||||
false : succ n <N a +N b → a +N (b +N c) <N succ n → False
|
||||
false pr1 pr2 rewrite Semiring.+Associative ℕSemiring a b c = cannotAddAndEnlarge' (<Transitive pr2 pr1)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inl (inl b+c<sn) | inl (inr sn<a+'b+c) with totality (subtractionNResult.result (-N (inl sn<a+b)) +N c) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inl (inl b+c<sn) | inl (inr sn<a+'b+c) | inl (inl x) = equalityZn _ _ (help4 {a} {b} {c} {succ n} sn<a+'b+c sn<a+b)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inl (inl b+c<sn) | inl (inr sn<a+'b+c) | inl (inr x) = exFalso (help18 {a} {b} {c} {succ n} b+c<sn sn<a+b aLess x)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inl (inl b+c<sn) | inl (inr sn<a+'b+c) | inr x = exFalso (help19 {a} {b} {c} {succ n} b+c<sn sn<a+b aLess x)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inl (inl b+c<sn) | inr a+'b+c=sn = exFalso (false sn<a+b a+'b+c=sn)
|
||||
where
|
||||
false : (succ n <N a +N b) → a +N (b +N c) ≡ succ n → False
|
||||
false pr1 pr2 rewrite Semiring.+Associative ℕSemiring a b c | equalityCommutative pr2 = cannotAddAndEnlarge' pr1
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inl (inr sn<b+c) with totality (a +N subtractionNResult.result (-N (inl sn<b+c))) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inl (inr sn<b+c) | inl (inl x) with totality (subtractionNResult.result (-N (inl sn<a+b)) +N c) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inl (inr sn<b+c) | inl (inl x) | inl (inl x₁) = equalityZn _ _ (help5 {a} {b} {c} {succ n} sn<b+c sn<a+b)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inl (inr sn<b+c) | inl (inl x) | inl (inr x1) = equalityZn _ _ (help16 {a} {b} {c} {succ n} sn<b+c sn<a+b x x1)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inl (inr sn<b+c) | inl (inl x) | inr x1 = exFalso (help17 {a} {b} {c} {succ n} sn<b+c sn<a+b x x1)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inl (inr sn<b+c) | inl (inr x) with totality (subtractionNResult.result (-N (inl sn<a+b)) +N c) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inl (inr sn<b+c) | inl (inr x) | inl (inl x1) = equalityZn _ _ (exFalso (help15 {a} {b} {c} {succ n} sn<b+c sn<a+b x x1))
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inl (inr sn<b+c) | inl (inr x) | inl (inr x1) = equalityZn _ _ (help14 {a} {b} {c} {succ n} sn<b+c sn<a+b x x1)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inl (inr sn<b+c) | inl (inr x) | inr x1 = equalityZn _ _ (exFalso (help13 {a} {b} {c} {succ n} sn<b+c sn<a+b x x1))
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inl (inr sn<b+c) | inr x with totality (subtractionNResult.result (-N (inl sn<a+b)) +N c) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inl (inr sn<b+c) | inr x | inl (inl x1) = equalityZn _ _ (exFalso (help12 {a} {b} {c} {succ n} sn<b+c sn<a+b x x1))
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inl (inr sn<b+c) | inr x | inl (inr x1) = equalityZn _ _ (exFalso (help10 {a} {b} {c} {succ n} sn<b+c sn<a+b x x1))
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inl (inr sn<b+c) | inr x | inr x₁ = equalityZn _ _ refl
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inr b+c=sn with totality (a +N 0) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inr b+c=sn | inl (inl a+0<sn) with totality (subtractionNResult.result (-N (inl sn<a+b)) +N c) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inr b+c=sn | inl (inl a+0<sn) | inl (inl x) = equalityZn _ _ (help6 {a} {b} {c} {succ n} b+c=sn sn<a+b)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inr b+c=sn | inl (inl _) | inl (inr x) = equalityZn _ _ (exFalso (help11 {a} {b} {c} {succ n} aLess b+c=sn sn<a+b x))
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inr b+c=sn | inl (inl a+0<sn) | inr x = equalityZn _ _ (exFalso (help7 {a} {b} {c} {succ n} b+c=sn aLess sn<a+b x))
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inr b+c=sn | inl (inr sn<a+0) = exFalso (help8 {a} {succ n} sn<a+0 aLess)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inr b+c=sn | inr a+0=sn = exFalso (help9 {a} {succ n} a+0=sn aLess)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b with totality c (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inl (inl c<sn) with totality (b +N c) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inl (inl c<sn) | inl (inl b+c<sn) with totality (a +N (b +N c)) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inl (inl c<sn) | inl (inl b+c<sn) | inl (inl a+'b+c<sn) = equalityZn _ _ (help27 {a} {b} {c} {n} sn=a+b a+'b+c<sn)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inl (inl c<sn) | inl (inl b+c<sn) | inl (inr sn<a+'b+c) = equalityZn _ _ (help28 {a} {b} {c} {succ n} sn<a+'b+c sn=a+b)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inl (inl c<sn) | inl (inl b+c<sn) | inr a+'b+c=sn = equalityZn _ _ (help26 {a} {b} {c} {succ n} sn=a+b a+'b+c=sn)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inl (inl c<sn) | inl (inr sn<b+c) with totality (a +N subtractionNResult.result (-N (inl sn<b+c))) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inl (inl c<sn) | inl (inr sn<b+c) | inl (inl x) = equalityZn _ _ (help31 {a} {b} {c} {succ n} sn=a+b sn<b+c)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inl (inl c<sn) | inl (inr sn<b+c) | inl (inr x) = exFalso (help30 {a} {b} {c} {succ n} cLess sn=a+b sn<b+c x)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inl (inl c<sn) | inl (inr sn<b+c) | inr x = exFalso (help29 {a} {b} {c} {succ n} cLess sn<b+c x sn=a+b)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inl (inl c<sn) | inr b+c=sn with totality (a +N 0) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inl (inl c<sn) | inr b+c=sn | inl (inl a+0<sn) = equalityZn _ _ (help25 {a} {b} {c} {succ n} sn=a+b b+c=sn)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inl (inl c<sn) | inr b+c=sn | inl (inr sn<a+0) = exFalso (help24 {a} {succ n} aLess sn<a+0)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inl (inl c<sn) | inr b+c=sn | inr a+0=sn = exFalso (help23 {a} {succ n} aLess a+0=sn)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inl (inr sn<c) = exFalso (TotalOrder.irreflexive ℕTotalOrder (<Transitive sn<c cLess))
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inr c=sn with totality (b +N c) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inr c=sn | inl (inl b+c<sn) = exFalso (help b+c<sn)
|
||||
where
|
||||
help : (b +N c <N succ n) → False
|
||||
help b+c<sn rewrite equalityCommutative c=sn | Semiring.commutative ℕSemiring b c = cannotAddAndEnlarge' b+c<sn
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inr c=sn | inl (inr sn<b+c) with totality (a +N subtractionNResult.result (-N (inl sn<b+c))) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inr c=sn | inl (inr sn<b+c) | inl (inl x) = exFalso (help20 {a} {b} {c} {succ n} c=sn sn=a+b sn<b+c x)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inr c=sn | inl (inr sn<b+c) | inl (inr x) = exFalso (help22 {a} {b} {c} {succ n} sn=a+b c=sn sn<b+c x)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inr c=sn | inl (inr sn<b+c) | inr x = equalityZn _ _ refl
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inr c=sn | inr b+c=sn = exFalso (help21 {a} {b} {c} {succ n} sn=a+b b+c=sn c=sn aLess)
|
||||
|
||||
subLess : {a b : ℕ} → (0<a : 0 <N a) → (a<b : a <N b) → subtractionNResult.result (-N (inl a<b)) <N b
|
||||
subLess {zero} {b} 0<a a<b = exFalso (TotalOrder.irreflexive ℕTotalOrder 0<a)
|
||||
subLess {succ a} {b} _ a<b with -N (inl a<b)
|
||||
... | record { result = b-a ; pr = pr } = record { x = a ; proof = pr }
|
||||
|
||||
inverseZn : {n : ℕ} → {pr : 0 <N n} → (a : ℤn n pr) → Sg (ℤn n pr) (λ i → i +n a ≡ record { x = 0 ; xLess = pr } )
|
||||
inverseZn {zero} {()}
|
||||
inverseZn {succ n} {0<n} record { x = zero ; xLess = zeroLess } = record { x = zero ; xLess = zeroLess } , plusZnIdentityLeft _
|
||||
inverseZn {succ n} {0<n} record { x = (succ a) ; xLess = aLess } = ans , pr
|
||||
where
|
||||
ans = record { x = subtractionNResult.result (-N (inl (canRemoveSuccFrom<N aLess))) ; xLess = subLess (succIsPositive a) aLess }
|
||||
pr : ans +n record { x = (succ a) ; xLess = aLess } ≡ record { x = 0 ; xLess = 0<n }
|
||||
pr with totality (subtractionNResult.result (-N (inl (canRemoveSuccFrom<N aLess))) +N (succ a)) (succ n)
|
||||
... | inl (inl x) = exFalso f
|
||||
where
|
||||
h : subtractionNResult.result (-N (inl (canRemoveSuccFrom<N aLess))) +N succ a ≡ succ n
|
||||
h with -N (inl (canRemoveSuccFrom<N aLess))
|
||||
h | record { result = result ; pr = pr } rewrite equalityCommutative pr = Semiring.commutative ℕSemiring result (succ a)
|
||||
f : False
|
||||
f = TotalOrder.irreflexive ℕTotalOrder (identityOfIndiscernablesLeft _<N_ x h)
|
||||
... | inl (inr x) = exFalso f
|
||||
where
|
||||
h : subtractionNResult.result (-N (inl (canRemoveSuccFrom<N aLess))) +N succ a ≡ succ n
|
||||
h = addMinus {succ a} {succ n} (inl aLess)
|
||||
f : False
|
||||
f rewrite h = TotalOrder.irreflexive ℕTotalOrder x
|
||||
... | (inr n-a+sa=sn) = equalityZn _ _ refl
|
||||
|
||||
ℤnGroup : (n : ℕ) → (pr : 0 <N n) → Group (reflSetoid (ℤn n pr)) _+n_
|
||||
ℤnGroup n pr = record { 0G = record { x = 0 ; xLess = pr } ; inverse = λ a → underlying (inverseZn a) ; +Associative = λ {a} {b} {c} → plusZnAssociative a b c ; identRight = λ {a} → plusZnIdentityRight a ; identLeft = λ {a} → plusZnIdentityLeft a ; invLeft = λ {a} → helpInvLeft a ; invRight = λ {a} → helpInvRight a ; +WellDefined = reflGroupWellDefined }
|
||||
where
|
||||
helpInvLeft : (a : ℤn n pr) → underlying (inverseZn a) +n a ≡ record { x = 0 ; xLess = pr }
|
||||
helpInvLeft a with inverseZn a
|
||||
... | -a , pr-a = pr-a
|
||||
helpInvRight : (a : ℤn n pr) → a +n underlying (inverseZn a) ≡ record { x = 0 ; xLess = pr }
|
||||
helpInvRight a rewrite plusZnCommutative a (underlying (inverseZn a)) = helpInvLeft a
|
||||
ℤnGroup : (n : ℕ) → .(pr : 0 <N n) → Group (reflSetoid (ℤn n pr)) (_+n_ pr)
|
||||
Group.+WellDefined (ℤnGroup n 0<n) refl refl = refl
|
||||
Group.0G (ℤnGroup n 0<n) = record { x = 0 ; xLess = 0<n }
|
||||
Group.inverse (ℤnGroup n 0<n) = inverseN 0<n
|
||||
Group.+Associative (ℤnGroup n 0<n) {a} {b} {c} = equalityCommutative (plusZnAssociative 0<n a b c)
|
||||
Group.identRight (ℤnGroup n 0<n) {a} = plusZnIdentityRight 0<n a
|
||||
Group.identLeft (ℤnGroup n 0<n) {a} = plusZnIdentityLeft 0<n a
|
||||
Group.invLeft (ℤnGroup n 0<n) {a} = invLeft 0<n a
|
||||
Group.invRight (ℤnGroup n 0<n) {a} = transitivity (plusZnCommutative 0<n a (inverseN 0<n a)) (invLeft 0<n a)
|
||||
|
||||
ℤnAbGroup : (n : ℕ) → (pr : 0 <N n) → AbelianGroup (ℤnGroup n pr)
|
||||
AbelianGroup.commutative (ℤnAbGroup n pr) {a} {b} = plusZnCommutative a b
|
||||
AbelianGroup.commutative (ℤnAbGroup n pr) {a} {b} = plusZnCommutative pr a b
|
||||
|
||||
ℤnFinite : (n : ℕ) → (pr : 0 <N n) → FiniteGroup (ℤnGroup n pr) (FinSet n)
|
||||
SetoidToSet.project (FiniteGroup.toSet (ℤnFinite n pr)) record { x = x ; xLess = xLess } = ofNat x xLess
|
||||
SetoidToSet.wellDefined (FiniteGroup.toSet (ℤnFinite n pr)) x y x=y rewrite x=y = refl
|
||||
Surjection.property (SetoidToSet.surj (FiniteGroup.toSet (ℤnFinite n pr))) b = record { x = toNat b ; xLess = toNatSmaller b } , ofNatToNat b
|
||||
SetoidToSet.inj (FiniteGroup.toSet (ℤnFinite zero ())) x y x=y
|
||||
SetoidToSet.inj (FiniteGroup.toSet (ℤnFinite (succ n) pr)) record { x = x ; xLess = xLess } record { x = y ; xLess = yLess } x=y = equalityZn _ _ b
|
||||
where
|
||||
b : x ≡ y
|
||||
b = ofNatInjective x y xLess yLess x=y
|
||||
FiniteSet.size (FiniteGroup.finite (ℤnFinite n pr)) = n
|
||||
FiniteSet.mapping (FiniteGroup.finite (ℤnFinite n pr)) = id
|
||||
FiniteSet.bij (FiniteGroup.finite (ℤnFinite n pr)) = idIsBijective
|
||||
SetoidToSet.project (FiniteGroup.toSet (ℤnFinite (succ n) 0<n)) record { x = x ; xLess = xLess } = ofNat x xLess
|
||||
SetoidToSet.wellDefined (FiniteGroup.toSet (ℤnFinite (succ n) 0<n)) x y x=y rewrite x=y = refl
|
||||
SetoidToSet.surj (FiniteGroup.toSet (ℤnFinite (succ n) 0<n)) b = record { x = toNat b ; xLess = toNatSmaller b } , ofNatToNat b
|
||||
SetoidToSet.inj (FiniteGroup.toSet (ℤnFinite (succ n) 0<n)) record { x = x ; xLess = xLess } record { x = y ; xLess = yLess } eq = equalityZn (ofNatInjective x y xLess yLess eq)
|
||||
FiniteGroup.finite (ℤnFinite n pr) = record { size = n ; mapping = id ; bij = idIsBijective }
|
||||
|
101
Numbers/Modulo/ModuloFunction.agda
Normal file
101
Numbers/Modulo/ModuloFunction.agda
Normal file
@@ -0,0 +1,101 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.EuclideanAlgorithm
|
||||
open import Semirings.Definition
|
||||
open import Orders.Total.Definition
|
||||
|
||||
module Numbers.Modulo.ModuloFunction where
|
||||
|
||||
open TotalOrder ℕTotalOrder
|
||||
|
||||
private
|
||||
notBigger : (a : ℕ) {n : ℕ} → succ a +N n ≡ a → False
|
||||
notBigger (succ a) {n} pr = notBigger a {n} (succInjective pr)
|
||||
|
||||
notBigger' : (a : ℕ) {n : ℕ} → succ a +N n ≡ n → False
|
||||
notBigger' (succ a) {succ n} pr rewrite Semiring.commutative ℕSemiring a (succ n) | Semiring.commutative ℕSemiring n a = notBigger' _ (succInjective pr)
|
||||
|
||||
abstract
|
||||
mod : (n : ℕ) → .(pr : 0 <N n) → (a : ℕ) → ℕ
|
||||
mod (succ n) 0<n a = divisionAlgResult.rem (divisionAlg (succ n) a)
|
||||
|
||||
modIsMod : {n : ℕ} → .(pr : 0 <N n) → (a : ℕ) → a <N n → mod n pr a ≡ a
|
||||
modIsMod {succ n} 0<n a a<n with divisionAlg (succ n) a
|
||||
modIsMod {succ n} 0<n a a<n | record { quot = zero ; rem = rem ; pr = pr ; remIsSmall = remIsSmall ; quotSmall = quotSmall } rewrite multiplicationNIsCommutative n 0 = pr
|
||||
modIsMod {succ n} 0<n a (le x proof) | record { quot = succ quot ; rem = rem ; pr = pr ; remIsSmall = inl rem<sn ; quotSmall = quotSmall } = exFalso (notBigger a v)
|
||||
where
|
||||
t : ((succ (a +N x)) *N succ quot) +N rem ≡ a
|
||||
t rewrite Semiring.commutative ℕSemiring a x = transitivity (applyEquality (λ i → (i *N succ quot) +N rem) proof) pr
|
||||
u : (succ quot *N succ (a +N x)) +N rem ≡ a
|
||||
u = transitivity (applyEquality (_+N rem) (multiplicationNIsCommutative (succ quot) _)) t
|
||||
v : succ a +N ((x +N quot *N succ (a +N x)) +N rem) ≡ a
|
||||
v = transitivity (applyEquality succ (transitivity (Semiring.+Associative ℕSemiring a _ rem) (applyEquality (_+N rem) (Semiring.+Associative ℕSemiring a x _)))) u
|
||||
|
||||
mod<N : {n : ℕ} → .(pr : 0 <N n) → (a : ℕ) → (mod n pr a) <N n
|
||||
mod<N {succ n} pr a with divisionAlg (succ n) a
|
||||
mod<N {succ n} 0<n a | record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = inl x ; quotSmall = quotSmall } = x
|
||||
|
||||
modIdempotent : {n : ℕ} → .(pr : 0 <N n) → (a : ℕ) → mod n pr (mod n pr a) ≡ mod n pr a
|
||||
modIdempotent 0<n a = modIsMod 0<n (mod _ 0<n a) (mod<N 0<n a)
|
||||
|
||||
modAnd+n : {n : ℕ} → .(pr : 0 <N n) → (a : ℕ) → mod n pr (n +N a) ≡ mod n pr a
|
||||
modAnd+n {succ n} 0<sn a with divisionAlg (succ n) (succ n +N a)
|
||||
modAnd+n {succ n} 0<sn a | record { quot = 0 ; rem = rem ; pr = pr ; remIsSmall = inl remIsSmall ; quotSmall = inl q } rewrite multiplicationNIsCommutative n 0 | pr = exFalso (notBigger (succ n) (transitivity (applyEquality succ (transitivity (applyEquality succ (Semiring.+Associative ℕSemiring n a _)) (Semiring.commutative ℕSemiring (succ (n +N a)) _))) (_<N_.proof remIsSmall)))
|
||||
modAnd+n {succ n} 0<sn a | record { quot = succ quot ; rem = rem ; pr = pr ; remIsSmall = inl remIsSmall ; quotSmall = inl q } = modIsUnique (record { quot = quot ; rem = rem ; pr = t ; remIsSmall = inl remIsSmall ; quotSmall = inl q }) (divisionAlg (succ n) a)
|
||||
where
|
||||
g : succ n +N ((quot *N succ n) +N rem) ≡ succ n +N a
|
||||
g = transitivity (Semiring.+Associative ℕSemiring (succ n) _ rem) (transitivity (applyEquality (_+N rem) (multiplicationNIsCommutative (succ quot) _)) pr)
|
||||
h : (quot *N succ n) +N rem ≡ a
|
||||
h = canSubtractFromEqualityLeft g
|
||||
t : (succ n *N quot) +N rem ≡ a
|
||||
t = transitivity (applyEquality (_+N rem) (multiplicationNIsCommutative (succ n) quot)) h
|
||||
|
||||
modExtracts : {n : ℕ} → .(pr : 0 <N n) → (a b : ℕ) → mod n pr (a +N b) ≡ mod n pr (mod n pr a +N mod n pr b)
|
||||
modExtracts {succ n} 0<sn a b with divisionAlg (succ n) a
|
||||
modExtracts {succ n} 0<sn a b | record { quot = sn/a ; rem = sn%a ; pr = prA ; remIsSmall = remIsSmallA ; quotSmall = quotSmallA } with divisionAlg (succ n) b
|
||||
modExtracts {succ n} 0<sn a b | record { quot = sn/a ; rem = sn%a ; pr = prA ; remIsSmall = remIsSmallA ; quotSmall = quotSmallA } | record { quot = sn/b ; rem = sn%b ; pr = prB ; remIsSmall = remIsSmallB ; quotSmall = quotSmallB } with divisionAlg (succ n) (a +N b)
|
||||
modExtracts {succ n} 0<sn a b | record { quot = sn/a ; rem = sn%a ; pr = prA ; remIsSmall = remIsSmallA ; quotSmall = quotSmallA } | record { quot = sn/b ; rem = sn%b ; pr = prB ; remIsSmall = remIsSmallB ; quotSmall = quotSmallB } | record { quot = sn/a+b ; rem = sn%a+b ; pr = prA+B ; remIsSmall = remIsSmallA+B ; quotSmall = quotSmallA+B } with divisionAlg (succ n) (sn%a +N sn%b)
|
||||
modExtracts {succ n} 0<sn a b | record { quot = sn/a ; rem = sn%a ; pr = prA ; remIsSmall = inl sn%a<sn ; quotSmall = inl _ } | record { quot = sn/b ; rem = sn%b ; pr = prB ; remIsSmall = inl sn%b<sn ; quotSmall = inl _ } | record { quot = sn/a+b ; rem = sn%a+b ; pr = prA+B ; remIsSmall = inl sn%a+b<sn ; quotSmall = inl _ } | record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = inl remIsSmall ; quotSmall = inl _ } = equalityCommutative whoa
|
||||
where
|
||||
t : ((succ n *N sn/a) +N sn%a) +N ((succ n *N sn/b) +N sn%b) ≡ a +N b
|
||||
t = +NWellDefined prA prB
|
||||
t' : (succ n *N (sn/a +N sn/b)) +N (sn%a +N sn%b) ≡ a +N b
|
||||
t' = transitivity (transitivity (Semiring.+Associative ℕSemiring (succ n *N (sn/a +N sn/b)) sn%a sn%b) (transitivity (applyEquality (_+N sn%b) (transitivity (transitivity (applyEquality (_+N sn%a) (transitivity (applyEquality (succ n *N_) (Semiring.commutative ℕSemiring sn/a sn/b)) (Semiring.+DistributesOver* ℕSemiring (succ n) sn/b sn/a))) (equalityCommutative (Semiring.+Associative ℕSemiring (succ n *N sn/b) (succ n *N sn/a) sn%a))) (Semiring.commutative ℕSemiring (succ n *N sn/b) ((succ n *N sn/a) +N sn%a)))) (equalityCommutative (Semiring.+Associative ℕSemiring ((succ n *N sn/a) +N sn%a) (succ n *N sn/b) sn%b)))) t
|
||||
thing : ((succ n) *N quot) +N rem ≡ sn%a +N sn%b
|
||||
thing = pr
|
||||
t'' : (succ n *N ((sn/a +N sn/b) +N quot)) +N rem ≡ a +N b
|
||||
t'' rewrite Semiring.+DistributesOver* ℕSemiring (succ n) (sn/a +N sn/b) quot | equalityCommutative (Semiring.+Associative ℕSemiring (succ n *N (sn/a +N sn/b)) (succ n *N quot) rem) | thing = t'
|
||||
whoa : rem ≡ sn%a+b
|
||||
whoa = modUniqueLemma _ _ remIsSmall sn%a+b<sn (transitivity t'' (equalityCommutative prA+B))
|
||||
|
||||
modAddition : {n : ℕ} → .(pr : 0 <N n) → {a b : ℕ} → (a <N n) → (b <N n) → (a +N b ≡ mod n pr (a +N b)) || (a +N b ≡ n +N mod n pr (a +N b))
|
||||
modAddition {succ n} 0<sn {a} {b} a<n b<n with divisionAlg (succ n) (a +N b)
|
||||
modAddition {succ n} 0<sn {a} {b} a<n b<n | record { quot = zero ; rem = rem ; pr = pr ; remIsSmall = inl rem<sn ; quotSmall = inl _ } = inl (equalityCommutative (transitivity (applyEquality (_+N rem) (multiplicationNIsCommutative 0 n)) pr))
|
||||
modAddition {succ n} 0<sn {a} {b} a<n b<n | record { quot = succ zero ; rem = rem ; pr = pr ; remIsSmall = inl rem<sn ; quotSmall = inl _ } = inr (transitivity (equalityCommutative pr) (applyEquality (λ i → succ i +N rem) (transitivity (multiplicationNIsCommutative n 1) (Semiring.sumZeroRight ℕSemiring n))))
|
||||
modAddition {succ n} 0<sn {a} {b} (le x prA) (le y prB) | record { quot = succ (succ quot) ; rem = 0 ; pr = pr ; remIsSmall = inl rem<sn ; quotSmall = inl _ } rewrite Semiring.sumZeroRight ℕSemiring (succ n *N succ (succ quot)) | multiplicationNIsCommutative (succ n) (succ (succ quot)) = exFalso (notBigger' ((x +N succ y) +N (quot *N succ n)) u)
|
||||
where
|
||||
t : ((succ x +N a) +N succ (y +N b)) +N (quot *N succ n) ≡ a +N b
|
||||
t = transitivity (transitivity (applyEquality (_+N quot *N succ n) (+NWellDefined prA prB)) (equalityCommutative (Semiring.+Associative ℕSemiring (succ n) (succ n) (quot *N succ n)))) pr
|
||||
u : ((succ x +N succ y) +N (quot *N succ n)) +N (a +N b) ≡ a +N b
|
||||
u rewrite Semiring.commutative ℕSemiring ((x +N succ y) +N quot *N succ n) (a +N b) | Semiring.+Associative ℕSemiring (succ a +N b) (x +N succ y) (quot *N succ n) = transitivity (applyEquality (λ i → succ i +N quot *N succ n) (transitivity (Semiring.commutative ℕSemiring (a +N b) _) (transitivity (equalityCommutative (Semiring.+Associative ℕSemiring x (succ y) _)) (transitivity (applyEquality (x +N_) (transitivity (Semiring.+Associative ℕSemiring (succ y) a b) (transitivity (applyEquality (_+N b) (Semiring.commutative ℕSemiring _ a)) (equalityCommutative (Semiring.+Associative ℕSemiring a (succ y) b))))) (Semiring.+Associative ℕSemiring x a _))))) t
|
||||
modAddition {succ n} 0<sn {a} {b} a<n b<n | record { quot = succ (succ quot) ; rem = succ rem ; pr = pr ; remIsSmall = inl rem<sn ; quotSmall = inl _ } = exFalso (irreflexive (<Transitive u t))
|
||||
where
|
||||
t : (succ n *N succ (succ quot) +N succ rem) <N succ n +N succ n
|
||||
t = identityOfIndiscernablesLeft _<N_ (addStrongInequalities a<n b<n) (equalityCommutative pr)
|
||||
lemm : {a c d : ℕ} → (a +N a) <N (a *N succ (succ c)) +N succ d
|
||||
lemm {a} {c} {d} = le (a *N c +N d) f
|
||||
where
|
||||
f : succ ((a *N c +N d) +N (a +N a)) ≡ a *N succ (succ c) +N succ d
|
||||
f rewrite multiplicationNIsCommutative a (succ (succ c)) | Semiring.+Associative ℕSemiring a a (c *N a) | multiplicationNIsCommutative c a | Semiring.commutative ℕSemiring (a *N c +N d) (a +N a) | Semiring.+Associative ℕSemiring (a +N a) (a *N c) d | Semiring.commutative ℕSemiring ((a +N a) +N a *N c) (succ d) = applyEquality succ (Semiring.commutative ℕSemiring _ d)
|
||||
u : (succ n +N succ n) <N (succ n *N succ (succ quot) +N succ rem)
|
||||
u = lemm {succ n}
|
||||
|
||||
modN : {n : ℕ} → .(0<n : 0 <N n) → mod n 0<n n ≡ 0
|
||||
modN {succ n} 0<n = modIsUnique (divisionAlg (succ n) (succ n)) record { quot = 1 ; rem = zero ; pr = ans ; remIsSmall = inl (le n (Semiring.sumZeroRight ℕSemiring (succ n))) ; quotSmall = inl (le n (Semiring.sumZeroRight ℕSemiring (succ n))) }
|
||||
where
|
||||
ans : succ (n *N 1 +N 0) ≡ succ n
|
||||
ans rewrite multiplicationNIsCommutative n 1 | Semiring.sumZeroRight ℕSemiring n | Semiring.sumZeroRight ℕSemiring n = refl
|
Reference in New Issue
Block a user