mirror of
https://github.com/Smaug123/agdaproofs
synced 2025-10-11 06:38:39 +00:00
Rearrange some of Naturals (#48)
This commit is contained in:
@@ -65,7 +65,7 @@ module Numbers.Modulo.IntegersModN where
|
||||
plusZnIdentityLeft {zero} {()}
|
||||
plusZnIdentityLeft {succ n} {pr} record { x = x ; xLess = xLess } with orderIsTotal 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 (cannotBeLeqAndG {x} {succ n} (inl xLess) succn<x)
|
||||
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)
|
||||
|
||||
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)))
|
||||
@@ -203,7 +203,7 @@ module Numbers.Modulo.IntegersModN where
|
||||
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 = naughtE (equalityCommutative b+c=c)
|
||||
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
|
||||
@@ -590,7 +590,7 @@ module Numbers.Modulo.IntegersModN where
|
||||
where
|
||||
a=0 : (a : ℕ) → (a +N a ≡ a) → a ≡ 0
|
||||
a=0 zero pr = refl
|
||||
a=0 (succ a) pr = naughtE {_} {a} (equalityCommutative (canSubtractFromEqualityRight pr))
|
||||
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 orderIsTotal (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 orderIsTotal (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)
|
||||
@@ -661,18 +661,19 @@ module Numbers.Modulo.IntegersModN where
|
||||
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 (subtractOneFromInequality aLess))) ; xLess = subLess (succIsPositive a) aLess }
|
||||
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 orderIsTotal (subtractionNResult.result (-N (inl (subtractOneFromInequality aLess))) +N (succ a)) (succ n)
|
||||
pr with orderIsTotal (subtractionNResult.result (-N (inl (canRemoveSuccFrom<N aLess))) +N (succ a)) (succ n)
|
||||
... | inl (inl x) = exFalso f
|
||||
where
|
||||
h : subtractionNResult.result (-N (inl (le (_<N_.x aLess) (transitivity (equalityCommutative (succExtracts (_<N_.x aLess) a)) (succInjective (_<N_.proof aLess)))))) +N succ a ≡ succ n
|
||||
h = addMinus {succ a} {succ n} (inl aLess)
|
||||
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 rewrite h = TotalOrder.irreflexive ℕTotalOrder x
|
||||
f = TotalOrder.irreflexive ℕTotalOrder (identityOfIndiscernablesLeft _ _ _ _<N_ x h)
|
||||
... | inl (inr x) = exFalso f
|
||||
where
|
||||
h : subtractionNResult.result (-N (inl (subtractOneFromInequality aLess))) +N succ a ≡ succ n
|
||||
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
|
||||
|
@@ -10,9 +10,6 @@ _+N_ : ℕ → ℕ → ℕ
|
||||
zero +N y = y
|
||||
succ x +N y = succ (x +N y)
|
||||
|
||||
addzeroLeft : (x : ℕ) → (zero +N x) ≡ x
|
||||
addzeroLeft x = refl
|
||||
|
||||
addZeroRight : (x : ℕ) → (x +N zero) ≡ x
|
||||
addZeroRight zero = refl
|
||||
addZeroRight (succ x) rewrite addZeroRight x = refl
|
||||
@@ -25,7 +22,7 @@ succCanMove : (x y : ℕ) → (x +N succ y) ≡ (succ x +N y)
|
||||
succCanMove x y = transitivity (succExtracts x y) refl
|
||||
|
||||
additionNIsCommutative : (x y : ℕ) → (x +N y) ≡ (y +N x)
|
||||
additionNIsCommutative zero y = transitivity (addzeroLeft y) (equalityCommutative (addZeroRight y))
|
||||
additionNIsCommutative zero y = equalityCommutative (addZeroRight y)
|
||||
additionNIsCommutative (succ x) zero = transitivity (addZeroRight (succ x)) refl
|
||||
additionNIsCommutative (succ x) (succ y) = transitivity refl (applyEquality succ (transitivity (succCanMove x y) (additionNIsCommutative (succ x) y)))
|
||||
|
||||
@@ -41,3 +38,10 @@ additionNIsAssociative (succ a) (succ b) c = transitivity refl (transitivity ref
|
||||
|
||||
succIsAddOne : (a : ℕ) → succ a ≡ a +N succ zero
|
||||
succIsAddOne a = equalityCommutative (transitivity (additionNIsCommutative a (succ zero)) refl)
|
||||
|
||||
canSubtractFromEqualityRight : {a b c : ℕ} → (a +N b ≡ c +N b) → a ≡ c
|
||||
canSubtractFromEqualityRight {a} {zero} {c} pr = transitivity (equalityCommutative (addZeroRight a)) (transitivity pr (addZeroRight c))
|
||||
canSubtractFromEqualityRight {a} {succ b} {c} pr rewrite additionNIsCommutative a (succ b) | additionNIsCommutative c (succ b) | additionNIsCommutative b a | additionNIsCommutative b c = canSubtractFromEqualityRight {a} {b} {c} (succInjective pr)
|
||||
|
||||
canSubtractFromEqualityLeft : {a b c : ℕ} → (a +N b ≡ a +N c) → b ≡ c
|
||||
canSubtractFromEqualityLeft {a} {b} {c} pr rewrite additionNIsCommutative a b | additionNIsCommutative a c = canSubtractFromEqualityRight {b} {a} {c} pr
|
||||
|
@@ -14,3 +14,10 @@ infix 100 succ
|
||||
|
||||
succInjective : {a b : ℕ} → (succ a ≡ succ b) → a ≡ b
|
||||
succInjective {a} {.a} refl = refl
|
||||
|
||||
naughtE : {a : ℕ} → zero ≡ succ a → False
|
||||
naughtE ()
|
||||
|
||||
aIsNotSuccA : (a : ℕ) → (a ≡ succ a) → False
|
||||
aIsNotSuccA zero pr = naughtE pr
|
||||
aIsNotSuccA (succ a) pr = aIsNotSuccA a (succInjective pr)
|
||||
|
22
Numbers/Naturals/Exponentiation.agda
Normal file
22
Numbers/Naturals/Exponentiation.agda
Normal file
@@ -0,0 +1,22 @@
|
||||
{-# OPTIONS --warning=error --safe --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Definition
|
||||
open import Numbers.Naturals.Addition
|
||||
open import Numbers.Naturals.Multiplication
|
||||
open import Numbers.Naturals.Order
|
||||
|
||||
module Numbers.Naturals.Exponentiation where
|
||||
|
||||
_^N_ : ℕ → ℕ → ℕ
|
||||
a ^N zero = 1
|
||||
a ^N succ b = a *N (a ^N b)
|
||||
|
||||
exponentiationIncreases : (a b : ℕ) → (a ≡ 0) || (a ≤N a ^N (succ b))
|
||||
exponentiationIncreases zero b = inl refl
|
||||
exponentiationIncreases (succ a) zero = inr (inr (applyEquality succ (transitivity (additionNIsCommutative 0 a) (multiplicationNIsCommutative 1 a))))
|
||||
exponentiationIncreases (succ a) (succ b) with exponentiationIncreases (succ a) b
|
||||
exponentiationIncreases (succ a) (succ b) | inr (inl x) = inr (inl (canAddToOneSideOfInequality _ x))
|
||||
exponentiationIncreases (succ a) (succ b) | inr (inr x) with productOne x
|
||||
exponentiationIncreases (succ 0) (succ b) | inr (inr x) | inr pr rewrite pr = inr (inr refl)
|
||||
exponentiationIncreases (succ (succ a)) (succ b) | inr (inr x) | inr pr rewrite pr | productWithOneRight a = inr (inl (le (succ (a +N a *N succ (succ a))) (additionNIsCommutative _ (succ (succ a)))))
|
@@ -80,3 +80,15 @@ multiplicationNIsAssociative (succ a) b c =
|
||||
transitivity refl
|
||||
(transitivity refl
|
||||
(transitivity (applyEquality ((λ x → b *N c +N x)) (multiplicationNIsAssociative a b c)) (transitivity (equalityCommutative (productDistributesRight b (a *N b) c)) refl)))
|
||||
|
||||
productOne : {a b : ℕ} → a ≡ a *N b → (a ≡ 0) || (b ≡ 1)
|
||||
productOne {zero} {b} pr = inl refl
|
||||
productOne {succ a} {zero} pr rewrite multiplicationNIsCommutative a 0 = exFalso (naughtE (equalityCommutative pr))
|
||||
productOne {succ a} {succ zero} pr = inr refl
|
||||
productOne {succ a} {succ (succ b)} pr rewrite multiplicationNIsCommutative a (succ (succ b)) | additionNIsCommutative (succ b) (a +N (a +N b *N a)) | additionNIsAssociative (succ a) (a +N b *N a) (succ b) | additionNIsCommutative (a +N b *N a) (succ b) = exFalso (bad pr)
|
||||
where
|
||||
bad' : {x : ℕ} → x ≡ succ x → False
|
||||
bad' ()
|
||||
bad : {x y : ℕ} → x ≡ x +N succ y → False
|
||||
bad {succ x} {zero} pr rewrite additionNIsCommutative x 1 = bad' pr
|
||||
bad {succ x} {succ y} pr = bad (succInjective pr)
|
||||
|
@@ -9,15 +9,17 @@ open import Numbers.Naturals.Definition
|
||||
open import Numbers.Naturals.Addition
|
||||
open import Numbers.Naturals.Order
|
||||
open import Numbers.Naturals.Multiplication
|
||||
open import Numbers.Naturals.Exponentiation
|
||||
open import Semirings.Definition
|
||||
open import Monoids.Definition
|
||||
|
||||
module Numbers.Naturals.Naturals where
|
||||
|
||||
open Numbers.Naturals.Definition using (ℕ ; zero ; succ; succInjective) public
|
||||
open Numbers.Naturals.Addition using (_+N_) public
|
||||
open Numbers.Naturals.Definition using (ℕ ; zero ; succ ; succInjective ; naughtE) public
|
||||
open Numbers.Naturals.Addition using (_+N_ ; canSubtractFromEqualityRight ; canSubtractFromEqualityLeft) public
|
||||
open Numbers.Naturals.Multiplication using (_*N_ ; multiplicationNIsCommutative) public
|
||||
open Numbers.Naturals.Order using (_<N_ ; le; succPreservesInequality) public
|
||||
open Numbers.Naturals.Exponentiation using (_^N_) public
|
||||
open Numbers.Naturals.Order using (_<N_ ; le; succPreservesInequality; succIsPositive; addingIncreases; zeroNeverGreater; noIntegersBetweenXAndSuccX; a<SuccA; canRemoveSuccFrom<N) public
|
||||
|
||||
ℕSemiring : Semiring 0 1 _+N_ _*N_
|
||||
Monoid.associative (Semiring.monoid ℕSemiring) a b c = equalityCommutative (additionNIsAssociative a b c)
|
||||
@@ -31,69 +33,14 @@ module Numbers.Naturals.Naturals where
|
||||
Semiring.productZeroRight ℕSemiring a = multiplicationNIsCommutative a 0
|
||||
Semiring.+DistributesOver* ℕSemiring = productDistributes
|
||||
|
||||
canAddZeroOnLeft : {a b : ℕ} → (a <N b) → (a +N zero) <N b
|
||||
canAddZeroOnLeft {a} {b} prAB = identityOfIndiscernablesLeft a b (a +N zero) (λ x y -> x <N y) prAB (equalityCommutative (addZeroRight a))
|
||||
canAddZeroOnRight : {a b : ℕ} → (a <N b) → a <N (b +N zero)
|
||||
canAddZeroOnRight {a} {b} prAB = identityOfIndiscernablesRight a b (b +N zero) (λ x y → x <N y) prAB (equalityCommutative (addZeroRight b))
|
||||
|
||||
canRemoveZeroFromLeft : {a b : ℕ} → (a +N zero) <N b → (a <N b)
|
||||
canRemoveZeroFromLeft {a} {b} prAB = identityOfIndiscernablesLeft (a +N zero) b a (λ x y → x <N y) prAB (addZeroRight a)
|
||||
canRemoveZeroFromRight : {a b : ℕ} → (a <N b +N zero) → a <N b
|
||||
canRemoveZeroFromRight {a} {b} prAB = identityOfIndiscernablesRight a (b +N zero) b (λ x y → x <N y) prAB (addZeroRight b)
|
||||
|
||||
collapseSuccOnLeft : {a b c : ℕ} → (succ (a +N c) <N b) → (a +N succ c <N b)
|
||||
collapseSuccOnLeft {a} {b} {c} pr = identityOfIndiscernablesLeft (succ (a +N c)) b (a +N succ c) (λ x y → x <N y) pr (equalityCommutative (succExtracts a c))
|
||||
|
||||
extractSuccOnLeft : {a b c : ℕ} → (a +N succ c <N b) → (succ (a +N c) <N b)
|
||||
extractSuccOnLeft {a} {b} {c} pr = identityOfIndiscernablesLeft (a +N succ c) b (succ (a +N c)) (λ x y → x <N y) pr (succExtracts a c)
|
||||
|
||||
collapseSuccOnRight : {a b c : ℕ} → (a <N succ (b +N c)) → (a <N b +N succ c)
|
||||
collapseSuccOnRight {a} {b} {c} pr = identityOfIndiscernablesRight a (succ (b +N c)) (b +N succ c) (λ x y → x <N y) pr (equalityCommutative (succExtracts b c))
|
||||
|
||||
extractSuccOnRight : {a b c : ℕ} → (a <N b +N succ c) → (a <N succ (b +N c))
|
||||
extractSuccOnRight {a} {b} {c} pr = identityOfIndiscernablesRight a (b +N succ c) (succ (b +N c)) (λ x y → x <N y) pr (succExtracts b c)
|
||||
|
||||
subtractOneFromInequality : {a b : ℕ} → (succ a <N succ b) → (a <N b)
|
||||
subtractOneFromInequality {a} {b} (le x proof) = le x (transitivity t pr')
|
||||
where
|
||||
pr' : x +N succ a ≡ b
|
||||
pr' = succInjective proof
|
||||
t : succ (x +N a) ≡ x +N succ a
|
||||
t = equalityCommutative (succExtracts x a)
|
||||
|
||||
succIsPositive : (a : ℕ) → zero <N succ a
|
||||
succIsPositive a = logical<NImpliesLe (record {})
|
||||
|
||||
a<SuccA : (a : ℕ) → a <N succ a
|
||||
a<SuccA a = le zero refl
|
||||
|
||||
record subtractionNResult (a b : ℕ) (p : a ≤N b) : Set where
|
||||
field
|
||||
result : ℕ
|
||||
pr : a +N result ≡ b
|
||||
|
||||
canSubtractFromEqualityRight : {a b c : ℕ} → (a +N b ≡ c +N b) → a ≡ c
|
||||
canSubtractFromEqualityRight {a} {zero} {c} pr = transitivity (equalityCommutative (addZeroRight a)) (transitivity pr (addZeroRight c))
|
||||
canSubtractFromEqualityRight {a} {succ b} {c} pr = canSubtractFromEqualityRight {a} {b} {c} h
|
||||
where
|
||||
k : a +N succ b ≡ succ (c +N b)
|
||||
k = identityOfIndiscernablesRight (a +N succ b) (c +N succ b) (succ (c +N b)) _≡_ pr (succExtracts c b)
|
||||
i : succ (a +N b) ≡ succ (c +N b)
|
||||
i = identityOfIndiscernablesLeft (a +N succ b) (succ (c +N b)) (succ (a +N b)) _≡_ k (succExtracts a b)
|
||||
h : a +N b ≡ c +N b
|
||||
h = succInjective i
|
||||
|
||||
canSubtractFromEqualityLeft : {a b c : ℕ} → (a +N b ≡ a +N c) → b ≡ c
|
||||
canSubtractFromEqualityLeft {a} {b} {c} pr = canSubtractFromEqualityRight {b} {a} {c} h
|
||||
where
|
||||
i : a +N b ≡ c +N a
|
||||
i = identityOfIndiscernablesRight (a +N b) (a +N c) (c +N a) _≡_ pr (additionNIsCommutative a c)
|
||||
h : b +N a ≡ c +N a
|
||||
h = identityOfIndiscernablesLeft (a +N b) (c +N a) (b +N a) _≡_ i (additionNIsCommutative a b)
|
||||
|
||||
subtractionNWellDefined : {a b : ℕ} → (p1 p2 : a ≤N b) → (s : subtractionNResult a b p1) → (t : subtractionNResult a b p2) → (subtractionNResult.result s ≡ subtractionNResult.result t)
|
||||
subtractionNWellDefined {a} {b} (inl x) pr2 record { result = result1 ; pr = pr1 } record { result = result ; pr = pr } = canSubtractFromEqualityLeft {a} (transitivity pr1 (equalityCommutative pr))
|
||||
subtractionNWellDefined {a} {.a} (inr refl) pr2 record { result = result1 ; pr = pr1 } record { result = result2 ; pr = pr } = transitivity g' (equalityCommutative g)
|
||||
subtractionNWellDefined : {a b : ℕ} → {p1 p2 : a ≤N b} → (s : subtractionNResult a b p1) → (t : subtractionNResult a b p2) → (subtractionNResult.result s ≡ subtractionNResult.result t)
|
||||
subtractionNWellDefined {a} {b} {inl x} {pr2} record { result = result1 ; pr = pr1 } record { result = result ; pr = pr } = canSubtractFromEqualityLeft {a} (transitivity pr1 (equalityCommutative pr))
|
||||
subtractionNWellDefined {a} {.a} {inr refl} {pr2} record { result = result1 ; pr = pr1 } record { result = result2 ; pr = pr } = transitivity g' (equalityCommutative g)
|
||||
where
|
||||
g : result2 ≡ 0
|
||||
g = canSubtractFromEqualityLeft {a} {_} {0} (transitivity pr (equalityCommutative (addZeroRight a)))
|
||||
@@ -104,7 +51,7 @@ module Numbers.Naturals.Naturals where
|
||||
-N {zero} {b} prAB = record { result = b ; pr = refl }
|
||||
-N {succ a} {zero} (inl (le x ()))
|
||||
-N {succ a} {zero} (inr ())
|
||||
-N {succ a} {succ b} (inl x) with -N {a} {b} (inl (subtractOneFromInequality x))
|
||||
-N {succ a} {succ b} (inl x) with -N {a} {b} (inl (canRemoveSuccFrom<N x))
|
||||
-N {succ a} {succ b} (inl x) | record { result = result ; pr = pr } = record { result = result ; pr = applyEquality succ pr }
|
||||
-N {succ a} {succ .a} (inr refl) = record { result = 0 ; pr = applyEquality succ (addZeroRight a) }
|
||||
|
||||
@@ -121,7 +68,7 @@ module Numbers.Naturals.Naturals where
|
||||
addMinus {succ a} {zero} (inl (le x ()))
|
||||
addMinus {succ a} {zero} (inr ())
|
||||
addMinus {succ a} {succ b} (inl x) with (-N {succ a} {succ b} (inl x))
|
||||
addMinus {succ a} {succ b} (inl x) | record { result = result ; pr = pr } = g
|
||||
addMinus {succ a} {succ b} (inl x) | record { result = result ; pr = pr } = transitivity (transitivity (applyEquality (_+N succ a) (transitivity (subtractionNWellDefined {p1 = inl (canRemoveSuccFrom<N x)} (record { result = subtractionNResult.result (-N (inl (canRemoveSuccFrom<N x))) ; pr = transitivity (additionNIsCommutative a _) (addMinus (inl (canRemoveSuccFrom<N x)))}) previous) (equalityCommutative t))) (additionNIsCommutative result (succ a))) pr
|
||||
where
|
||||
pr'' : (a <N b) || (a ≡ b)
|
||||
pr'' = (inl (le (_<N_.x x) (transitivity (equalityCommutative (succExtracts (_<N_.x x) a)) (succInjective (_<N_.proof x)))))
|
||||
@@ -130,30 +77,22 @@ module Numbers.Naturals.Naturals where
|
||||
next : Sg (subtractionNResult (succ a) (succ b) (addOneToWeakInequality pr'')) λ n → subtractionNResult.result n ≡ subtractionNResult.result previous
|
||||
next = bumpUpSubtraction pr'' previous
|
||||
t : result ≡ subtractionNResult.result (underlying next)
|
||||
t = subtractionNWellDefined {succ a} {succ b} (inl x) (addOneToWeakInequality pr'') (record { result = result ; pr = pr }) (underlying next)
|
||||
u : subtractionNResult.result previous ≡ subtractionNResult.result (underlying next)
|
||||
u = refl
|
||||
v : subtractionNResult.result previous ≡ result
|
||||
v = transitivity u (equalityCommutative t)
|
||||
lemma : (a b : ℕ) → succ (a +N b) ≡ b +N succ a
|
||||
lemma a b rewrite additionNIsCommutative a b = equalityCommutative (succExtracts b a)
|
||||
g : subtractionNResult.result previous +N succ a ≡ succ b
|
||||
g rewrite v = identityOfIndiscernablesLeft (succ (a +N result)) (succ b) (result +N succ a) _≡_ pr (lemma a result)
|
||||
t = subtractionNWellDefined {succ a} {succ b} {inl x} {addOneToWeakInequality pr''} (record { result = result ; pr = pr }) (underlying next)
|
||||
addMinus {succ a} {succ .a} (inr refl) = refl
|
||||
|
||||
addMinus' : {a b : ℕ} → (pr : a ≤N b) → a +N subtractionNResult.result (-N {a} {b} pr) ≡ b
|
||||
addMinus' {a} {b} pr rewrite additionNIsCommutative a (subtractionNResult.result (-N {a} {b} pr)) = addMinus {a} {b} pr
|
||||
|
||||
additionPreservesInequality : {a b : ℕ} → (c : ℕ) → a <N b → a +N c <N b +N c
|
||||
additionPreservesInequality {a} {b} zero prAB = canAddZeroOnRight {a +N zero} {b} (canAddZeroOnLeft {a} {b} prAB)
|
||||
additionPreservesInequality {a} {b} (succ c) prAB = collapseSuccOnLeft {a} {b +N succ c} {c} (collapseSuccOnRight {succ (a +N c)} {b} {c} (succPreservesInequality {a +N c} {b +N c} (additionPreservesInequality {a} {b} c prAB)))
|
||||
additionPreservesInequality {a} {b} zero prAB rewrite additionNIsCommutative a 0 | additionNIsCommutative b 0 = prAB
|
||||
additionPreservesInequality {a} {b} (succ c) (le x proof) = le x (transitivity (equalityCommutative (additionNIsAssociative (succ x) a (succ c))) (applyEquality (_+N succ c) proof))
|
||||
|
||||
additionPreservesInequalityOnLeft : {a b : ℕ} → (c : ℕ) → a <N b → c +N a <N c +N b
|
||||
additionPreservesInequalityOnLeft {a} {b} c prAB = identityOfIndiscernablesRight (c +N a) (b +N c) (c +N b) (λ a b → a <N b) (identityOfIndiscernablesLeft (a +N c) (b +N c) (c +N a) (λ a b → a <N b) (additionPreservesInequality {a} {b} c prAB) (additionNIsCommutative a c)) (additionNIsCommutative b c)
|
||||
|
||||
subtractionPreservesInequality : {a b : ℕ} → (c : ℕ) → a +N c <N b +N c → a <N b
|
||||
subtractionPreservesInequality {a} {b} zero prABC = canRemoveZeroFromRight {a} {b} (canRemoveZeroFromLeft {a} {b +N zero} prABC)
|
||||
subtractionPreservesInequality {a} {b} (succ c) prABC = subtractionPreservesInequality {a} {b} c (subtractOneFromInequality {a +N c} {b +N c} (extractSuccOnLeft {a} {succ (b +N c)} {c} (extractSuccOnRight {a +N succ c} {b} {c} prABC)))
|
||||
subtractionPreservesInequality {a} {b} zero prABC rewrite additionNIsCommutative a 0 | additionNIsCommutative b 0 = prABC
|
||||
subtractionPreservesInequality {a} {b} (succ c) (le x proof) = le x (canSubtractFromEqualityRight {b = succ c} (transitivity (additionNIsAssociative (succ x) a (succ c)) proof))
|
||||
|
||||
productNonzeroIsNonzero : {a b : ℕ} → zero <N a → zero <N b → zero <N a *N b
|
||||
productNonzeroIsNonzero {zero} {b} prA prB = prA
|
||||
@@ -177,9 +116,6 @@ module Numbers.Naturals.Naturals where
|
||||
k : succ b <N succ b +N (succ a *N succ b)
|
||||
k = identityOfIndiscernablesLeft (zero +N succ b) (succ b +N (succ a *N succ b)) (succ b) (λ a b → a <N b) j refl
|
||||
|
||||
naughtE : {B : Set} {a : ℕ} → (pr : zero ≡ succ a) → B
|
||||
naughtE {a} ()
|
||||
|
||||
productCancelsRight : (a b c : ℕ) → (zero <N a) → (b *N a ≡ c *N a) → (b ≡ c)
|
||||
productCancelsRight a zero zero aPos eq = refl
|
||||
productCancelsRight zero zero (succ c) (le x ()) eq
|
||||
@@ -188,7 +124,7 @@ module Numbers.Naturals.Naturals where
|
||||
h : zero ≡ succ c *N succ a
|
||||
h = eq
|
||||
contr : zero ≡ succ c
|
||||
contr = naughtE h
|
||||
contr = exFalso (naughtE h)
|
||||
|
||||
productCancelsRight zero (succ b) zero (le x ()) eq
|
||||
productCancelsRight (succ a) (succ b) zero aPos eq = contr
|
||||
@@ -196,7 +132,7 @@ module Numbers.Naturals.Naturals where
|
||||
h : succ b *N succ a ≡ zero
|
||||
h = eq
|
||||
contr : succ b ≡ zero
|
||||
contr = naughtE (equalityCommutative h)
|
||||
contr = exFalso (naughtE (equalityCommutative h))
|
||||
|
||||
productCancelsRight zero (succ b) (succ c) (le x ()) eq
|
||||
|
||||
@@ -227,9 +163,9 @@ module Numbers.Naturals.Naturals where
|
||||
productCancelsLeft' zero b c pr = inl refl
|
||||
productCancelsLeft' (succ a) b c pr = inr (productCancelsLeft (succ a) b c (succIsPositive a) pr)
|
||||
|
||||
lessRespectsAddition : (a b c : ℕ) → (a <N b) → ((a +N c) <N (b +N c))
|
||||
lessRespectsAddition a b zero prAB = canAddZeroOnRight {a +N zero} {b} (canAddZeroOnLeft {a} {b} prAB)
|
||||
lessRespectsAddition a b (succ c) prAB = collapseSuccOnRight {a +N succ c} {b} {c} (collapseSuccOnLeft {a} {succ (b +N c)} {c} (succPreservesInequality {a +N c} {b +N c} (lessRespectsAddition a b c prAB)))
|
||||
lessRespectsAddition : {a b : ℕ} (c : ℕ) → (a <N b) → ((a +N c) <N (b +N c))
|
||||
lessRespectsAddition {a} {b} zero prAB rewrite additionNIsCommutative a 0 | additionNIsCommutative b 0 = prAB
|
||||
lessRespectsAddition {a} {b} (succ c) prAB rewrite additionNIsCommutative a (succ c) | additionNIsCommutative b (succ c) | additionNIsCommutative c a | additionNIsCommutative c b = succPreservesInequality (lessRespectsAddition c prAB)
|
||||
|
||||
canTimesOneOnLeft : {a b : ℕ} → (a <N b) → (a *N (succ zero)) <N b
|
||||
canTimesOneOnLeft {a} {b} prAB = identityOfIndiscernablesLeft a b (a *N (succ zero)) (λ x y → x <N y) prAB (equalityCommutative (productWithOneRight a))
|
||||
@@ -268,7 +204,7 @@ module Numbers.Naturals.Naturals where
|
||||
h : c *N a <N c *N b
|
||||
h = lessRespectsMultiplicationLeft a b c (logical<NImpliesLe (leImpliesLogical<N prAB)) cPos
|
||||
j : c *N a +N c <N c *N b +N c
|
||||
j = lessRespectsAddition (c *N a) (c *N b) c h
|
||||
j = lessRespectsAddition c h
|
||||
i : c *N succ a <N c *N b +N c
|
||||
i = identityOfIndiscernablesLeft (c *N a +N c) (c *N b +N c) (c *N succ a) _<N_ j (equalityCommutative (bumpDownOnRight a c))
|
||||
m : c *N succ a <N c *N succ b
|
||||
@@ -327,9 +263,6 @@ module Numbers.Naturals.Naturals where
|
||||
-NIsDecreasing {a} {b} prAB with (-N (inl prAB))
|
||||
-NIsDecreasing {a} {b} (le x proof) | record { result = result ; pr = pr } = record { x = a ; proof = pr }
|
||||
|
||||
zeroNeverGreater : {a : ℕ} → (a <N zero) → False
|
||||
zeroNeverGreater {a} (le x ())
|
||||
|
||||
equalityN : (a b : ℕ) → Sg Bool (λ truth → if truth then a ≡ b else Unit)
|
||||
equalityN zero zero = ( BoolTrue , refl )
|
||||
equalityN zero (succ b) = ( BoolFalse , record {} )
|
||||
@@ -391,16 +324,6 @@ module Numbers.Naturals.Naturals where
|
||||
q' : b <N c
|
||||
q' = cancelInequalityLeft {a} {b} {c} q
|
||||
|
||||
<NWellDefined : {a b : ℕ} → (p1 : a <N b) → (p2 : a <N b) → _<N_.x p1 ≡ _<N_.x p2
|
||||
<NWellDefined {a} {b} (le x proof) (le y proof1) = equalityCommutative r
|
||||
where
|
||||
p : succ (y +N a) ≡ succ (x +N a)
|
||||
p = transitivity proof1 (equalityCommutative proof)
|
||||
q : y +N a ≡ x +N a
|
||||
q = succInjective {y +N a} {x +N a} p
|
||||
r : y ≡ x
|
||||
r = canSubtractFromEqualityRight {_} {a} q
|
||||
|
||||
cannotAddAndEnlarge : (a b : ℕ) → a <N succ (a +N b)
|
||||
cannotAddAndEnlarge a b = le b (applyEquality succ (additionNIsCommutative b a))
|
||||
|
||||
@@ -460,20 +383,8 @@ module Numbers.Naturals.Naturals where
|
||||
p'' = canSubtractFromEqualityLeft {d} {result} {c} p'
|
||||
equivalentSubtraction (succ a) b zero d (le x ()) prdb eq
|
||||
equivalentSubtraction (succ a) b (succ c) d prac prdb eq with (-N (inl (canRemoveSuccFrom<N prac)))
|
||||
equivalentSubtraction (succ a) b (succ c) d prac prdb eq | record { result = c-a ; pr = prc-a } = go
|
||||
where
|
||||
d-b : ℕ
|
||||
d-b = subtractionNResult.result (-N (inl prdb))
|
||||
s : subtractionNResult (succ a) (succ c) (inl prac)
|
||||
s = record { result = c-a ; pr = applyEquality succ prc-a }
|
||||
t : subtractionNResult.result (-N {a} {c} (inl (canRemoveSuccFrom<N prac))) ≡ subtractionNResult.result (-N {d} {b} (inl prdb))
|
||||
t = equivalentSubtraction a b c d (canRemoveSuccFrom<N prac) prdb (succInjective eq)
|
||||
t' : subtractionNResult.result (-N {a} {c} (inl (canRemoveSuccFrom<N prac))) ≡ d-b
|
||||
t' = transitivity t refl
|
||||
r : subtractionNResult.result (-N (inl (canRemoveSuccFrom<N prac))) ≡ subtractionNResult.result (-N (inl (le (_<N_.x prac) (transitivity (equalityCommutative (succExtracts (_<N_.x prac) a)) (succInjective (_<N_.proof prac))))))
|
||||
r = subtractionNWellDefined {a} {c} (inl (canRemoveSuccFrom<N prac)) ((inl (le (_<N_.x prac) (transitivity (equalityCommutative (succExtracts (_<N_.x prac) a)) (succInjective (_<N_.proof prac)))))) (-N (inl (canRemoveSuccFrom<N prac))) ((-N (inl (le (_<N_.x prac) (transitivity (equalityCommutative (succExtracts (_<N_.x prac) a)) (succInjective (_<N_.proof prac)))))))
|
||||
go : subtractionNResult.result (-N (inl (le (_<N_.x prac) (transitivity (equalityCommutative (succExtracts (_<N_.x prac) a)) (succInjective (_<N_.proof prac)))))) ≡ subtractionNResult.result (-N (inl prdb))
|
||||
go rewrite (equalityCommutative r) = t'
|
||||
equivalentSubtraction (succ a) b (succ c) d prac prdb eq | record { result = c-a ; pr = prc-a } with -N (inl prdb)
|
||||
equivalentSubtraction (succ a) b (succ c) d prac prdb eq | record { result = c-a ; pr = prc-a } | record { result = result ; pr = pr } rewrite equalityCommutative prc-a | equalityCommutative pr | equalityCommutative (additionNIsAssociative a d result) | additionNIsCommutative (a +N c-a) d | equalityCommutative (additionNIsAssociative d a c-a) | additionNIsCommutative a d = equalityCommutative (canSubtractFromEqualityLeft eq)
|
||||
|
||||
leLemma : (b c : ℕ) → (b ≤N c) ≡ (b +N zero ≤N c +N zero)
|
||||
leLemma b c rewrite addZeroRight c = q
|
||||
@@ -493,10 +404,6 @@ module Numbers.Naturals.Naturals where
|
||||
subtractionCast' : {a b c : ℕ} → {pr : a ≤N b} → (eq : b ≡ c) → (p : subtractionNResult a b pr) → Sg (subtractionNResult a c (lessCast' pr eq)) (λ res → subtractionNResult.result p ≡ subtractionNResult.result res)
|
||||
subtractionCast' {a} {b} {c} {a<b} eq subt rewrite eq = (subt , refl)
|
||||
|
||||
addingIncreases : (a b : ℕ) → a <N a +N succ b
|
||||
addingIncreases zero b = succIsPositive b
|
||||
addingIncreases (succ a) b = succPreservesInequality (addingIncreases a b)
|
||||
|
||||
addToRightWeakInequality : (a : ℕ) → {b c : ℕ} → (pr : b ≤N c) → (b ≤N c +N a)
|
||||
addToRightWeakInequality zero {b} {c} (inl x) rewrite (addZeroRight c) = inl x
|
||||
addToRightWeakInequality (succ a) {b} {c} (inl x) = inl (orderIsTransitive x (addingIncreases c a))
|
||||
@@ -577,7 +484,7 @@ module Numbers.Naturals.Naturals where
|
||||
s = transitivity q r
|
||||
s' : subtractionNResult.result resbc +N 0 ≡ subtractionNResult.result (-N {b +N 0 *N b} {c +N 0 *N c} (inl (lessRespectsMultiplicationLeft b c 1 b<c aPos)))
|
||||
s' = identityOfIndiscernablesLeft (subtractionNResult.result resbc) _ (subtractionNResult.result resbc +N 0) _≡_ s (equalityCommutative (addZeroRight _))
|
||||
r = subtractionNWellDefined {b +N 0 *N b} {c +N 0 *N c} ((lessCast' (lessCast (inl b<c) (equalityCommutative (addZeroRight b))) (equalityCommutative (addZeroRight c)))) (inl (lessRespectsMultiplicationLeft b c 1 b<c aPos)) (underlying resbc'') (-N {b +N 0 *N b} {c +N 0 *N c} (inl (lessRespectsMultiplicationLeft b c 1 b<c aPos)))
|
||||
r = subtractionNWellDefined {b +N 0 *N b} {c +N 0 *N c} {(lessCast' (lessCast (inl b<c) (equalityCommutative (addZeroRight b))) (equalityCommutative (addZeroRight c)))} {inl (lessRespectsMultiplicationLeft b c 1 b<c aPos)} (underlying resbc'') (-N {b +N 0 *N b} {c +N 0 *N c} (inl (lessRespectsMultiplicationLeft b c 1 b<c aPos)))
|
||||
g (a , b) = b
|
||||
g' (a , b) = b
|
||||
resbc'' = subtractionCast' (equalityCommutative (addZeroRight c)) (underlying resbc')
|
||||
@@ -609,20 +516,12 @@ module Numbers.Naturals.Naturals where
|
||||
u' rewrite equalityCommutative u = lhs'
|
||||
lhs' rewrite t = refl
|
||||
u = addSubIntoSub (inl b<c) (inl (lessRespectsMultiplicationLeft b c (succ a) b<c (succIsPositive a)))
|
||||
v = subtractionNWellDefined {succ (succ a) *N b} {succ (succ a) *N c} (addWeakInequalities (inl b<c) (inl (lessRespectsMultiplicationLeft b c (succ a) b<c (succIsPositive a)))) (inl (lessRespectsMultiplicationLeft b c (succ (succ a)) b<c aPos)) (-N {b +N (succ a *N b)} {c +N (succ a *N c)} (addWeakInequalities (inl b<c) (inl (lessRespectsMultiplicationLeft b c (succ a) b<c (succIsPositive a))))) (-N {(succ (succ a)) *N b} {(succ (succ a)) *N c} (inl (lessRespectsMultiplicationLeft b c (succ (succ a)) b<c aPos)))
|
||||
v = subtractionNWellDefined {succ (succ a) *N b} {succ (succ a) *N c} {addWeakInequalities (inl b<c) (inl (lessRespectsMultiplicationLeft b c (succ a) b<c (succIsPositive a)))} {inl (lessRespectsMultiplicationLeft b c (succ (succ a)) b<c aPos)} (-N {b +N (succ a *N b)} {c +N (succ a *N c)} (addWeakInequalities (inl b<c) (inl (lessRespectsMultiplicationLeft b c (succ a) b<c (succIsPositive a))))) (-N {(succ (succ a)) *N b} {(succ (succ a)) *N c} (inl (lessRespectsMultiplicationLeft b c (succ (succ a)) b<c aPos)))
|
||||
|
||||
subtractProduct' : {a b c : ℕ} → (aPos : 0 <N a) → (b<c : b <N c) →
|
||||
(subtractionNResult.result (-N (inl b<c))) *N a ≡ subtractionNResult.result (-N {a *N b} {a *N c} (inl (lessRespectsMultiplicationLeft b c a b<c aPos)))
|
||||
subtractProduct' {a} aPos b<c = identityOfIndiscernablesLeft (a *N subtractionNResult.result (-N (inl b<c))) _ (subtractionNResult.result (-N (inl b<c)) *N a) _≡_ (subtractProduct aPos b<c) (multiplicationNIsCommutative a _)
|
||||
|
||||
noIntegersBetweenXAndSuccX : {a : ℕ} (x : ℕ) → (x <N a) → (a <N succ x) → False
|
||||
noIntegersBetweenXAndSuccX {zero} x x<a a<sx = zeroNeverGreater x<a
|
||||
noIntegersBetweenXAndSuccX {succ a} x (le y proof) (le z proof1) with succInjective proof1
|
||||
... | ah rewrite (equalityCommutative proof) | (succExtracts z (y +N x)) | equalityCommutative (additionNIsAssociative (succ z) y x) = naughtE (equalityCommutative absurd)
|
||||
where
|
||||
absurd : succ (z +N y) ≡ 0
|
||||
absurd = canSubtractFromEqualityRight {_} {x} {0} ah
|
||||
|
||||
equalityDecidable : (a b : ℕ) → (a ≡ b) || ((a ≡ b) → False)
|
||||
equalityDecidable zero zero = inl refl
|
||||
equalityDecidable zero (succ b) = inr naughtE
|
||||
@@ -631,20 +530,12 @@ module Numbers.Naturals.Naturals where
|
||||
equalityDecidable (succ a) (succ b) | inl x = inl (applyEquality succ x)
|
||||
equalityDecidable (succ a) (succ b) | inr x = inr (λ t → x (succInjective t))
|
||||
|
||||
cannotBeLeqAndG : {a b : ℕ} → a ≤N b → b <N a → False
|
||||
cannotBeLeqAndG {a} {b} (inl x) b<a = orderIsIrreflexive x b<a
|
||||
cannotBeLeqAndG {a} {b} (inr prf) b<a = lessImpliesNotEqual b<a (equalityCommutative prf)
|
||||
|
||||
cannotAddKeepingEquality : (a b : ℕ) → (a ≡ a +N succ b) → False
|
||||
cannotAddKeepingEquality zero zero pr = naughtE pr
|
||||
cannotAddKeepingEquality (succ a) zero pr = cannotAddKeepingEquality a zero (succInjective pr)
|
||||
cannotAddKeepingEquality zero (succ b) pr = naughtE pr
|
||||
cannotAddKeepingEquality (succ a) (succ b) pr = cannotAddKeepingEquality a (succ b) (succInjective pr)
|
||||
|
||||
aIsNotSuccA : (a : ℕ) → (a ≡ succ a) → False
|
||||
aIsNotSuccA zero pr = naughtE pr
|
||||
aIsNotSuccA (succ a) pr = aIsNotSuccA a (succInjective pr)
|
||||
|
||||
ℕTotalOrder : TotalOrder ℕ
|
||||
PartialOrder._<_ (TotalOrder.order ℕTotalOrder) = _<N_
|
||||
PartialOrder.irreflexive (TotalOrder.order ℕTotalOrder) = lessIrreflexive
|
||||
|
@@ -24,6 +24,9 @@ infix 5 _≤N_
|
||||
_≤N_ : ℕ → ℕ → Set
|
||||
a ≤N b = (a <N b) || (a ≡ b)
|
||||
|
||||
succIsPositive : (a : ℕ) → zero <N succ a
|
||||
succIsPositive a = le a (applyEquality succ (additionNIsCommutative a 0))
|
||||
|
||||
aLessSucc : (a : ℕ) → (a <NLogical succ a)
|
||||
aLessSucc zero = record {}
|
||||
aLessSucc (succ a) = aLessSucc a
|
||||
@@ -70,4 +73,30 @@ succPreservesInequality : {a b : ℕ} → (a <N b) → (succ a <N succ b)
|
||||
succPreservesInequality {a} {b} prAB = logical<NImpliesLe (succPreservesInequalityLogical {a} {b} (leImpliesLogical<N prAB))
|
||||
|
||||
canRemoveSuccFrom<N : {a b : ℕ} → (succ a <N succ b) → (a <N b)
|
||||
canRemoveSuccFrom<N {a} {b} prAB = logical<NImpliesLe (leImpliesLogical<N prAB)
|
||||
canRemoveSuccFrom<N {a} {b} (le x proof) rewrite additionNIsCommutative x (succ a) | additionNIsCommutative a x = le x (succInjective proof)
|
||||
|
||||
a<SuccA : (a : ℕ) → a <N succ a
|
||||
a<SuccA a = le zero refl
|
||||
|
||||
canAddToOneSideOfInequality : {a b : ℕ} (c : ℕ) → a <N b → a <N (b +N c)
|
||||
canAddToOneSideOfInequality {a} {b} c (le x proof) = le (x +N c) (transitivity (applyEquality succ (additionNIsAssociative x c a)) (transitivity (applyEquality (λ i → (succ x) +N i) (additionNIsCommutative c a)) (transitivity (applyEquality succ (equalityCommutative (additionNIsAssociative x a c))) (applyEquality (_+N c) proof))))
|
||||
|
||||
addingIncreases : (a b : ℕ) → a <N a +N succ b
|
||||
addingIncreases zero b = succIsPositive b
|
||||
addingIncreases (succ a) b = succPreservesInequality (addingIncreases a b)
|
||||
|
||||
zeroNeverGreater : {a : ℕ} → (a <N zero) → False
|
||||
zeroNeverGreater {a} (le x ())
|
||||
|
||||
noIntegersBetweenXAndSuccX : {a : ℕ} (x : ℕ) → (x <N a) → (a <N succ x) → False
|
||||
noIntegersBetweenXAndSuccX {zero} x x<a a<sx = zeroNeverGreater x<a
|
||||
noIntegersBetweenXAndSuccX {succ a} x (le y proof) (le z proof1) with succInjective proof1
|
||||
... | ah rewrite (equalityCommutative proof) | (succExtracts z (y +N x)) | equalityCommutative (additionNIsAssociative (succ z) y x) | additionNIsCommutative (succ (z +N y)) x = lessIrreflexive {x} (le (z +N y) (transitivity (additionNIsCommutative _ x) ah))
|
||||
|
||||
<NWellDefined : {a b : ℕ} → (p1 : a <N b) → (p2 : a <N b) → _<N_.x p1 ≡ _<N_.x p2
|
||||
<NWellDefined {a} {b} (le x proof) (le y proof1) = equalityCommutative r
|
||||
where
|
||||
q : y +N a ≡ x +N a
|
||||
q = succInjective {y +N a} {x +N a} (transitivity proof1 (equalityCommutative proof))
|
||||
r : y ≡ x
|
||||
r = canSubtractFromEqualityRight q
|
||||
|
@@ -96,7 +96,7 @@ module Numbers.Primes.IntegerFactorisation where
|
||||
factorIntegerLemma zero indHyp = inl (succIsPositive 1)
|
||||
factorIntegerLemma (succ zero) indHyp = inl (succPreservesInequality (succIsPositive 0))
|
||||
factorIntegerLemma (succ (succ x)) indHyp with everyNumberHasAPrimeFactor {succ (succ x)} (succPreservesInequality (succIsPositive x))
|
||||
factorIntegerLemma (succ (succ x)) indHyp | a , record { fst = record { fst = divides record {quot = zero ; rem = .0 ; pr = ssxDivA ; remIsSmall = r } refl ; snd = 1<a }; snd = record { fst = primeA ; snd = smallerFactors } } rewrite Semiring.sumZeroRight ℕSemiring (a *N zero) | multiplicationNIsCommutative a 0 = naughtE ssxDivA
|
||||
factorIntegerLemma (succ (succ x)) indHyp | a , record { fst = record { fst = divides record {quot = zero ; rem = .0 ; pr = ssxDivA ; remIsSmall = r } refl ; snd = 1<a }; snd = record { fst = primeA ; snd = smallerFactors } } rewrite Semiring.sumZeroRight ℕSemiring (a *N zero) | multiplicationNIsCommutative a 0 = exFalso (naughtE ssxDivA)
|
||||
factorIntegerLemma (succ (succ x)) indHyp | a , record { fst = record { fst = divides record {quot = succ zero ; rem = .0 ; pr = ssxDivA ; remIsSmall = r } refl ; snd = 1<a }; snd = record { fst = primeA ; snd = smallerFactors } } = inr record { 1<a = succPreservesInequality (succIsPositive x) ; firstFactor = a ; firstFactorNontrivial = Prime.p>1 primeA ; firstFactorBiggerMin = inl (Prime.p>1 primeA) ; firstFactorDivision = record { quot = 1 ; rem = 0 ; pr = ssxDivA ; remIsSmall = r ; quotSmall = inl (TotalOrder.transitive ℕTotalOrder (le zero refl) 1<a) } ; firstFactorDivides = refl ; firstFactorPrime = record { p>1 = Prime.p>1 primeA ; pr = Prime.pr primeA } ; otherFactors = inl record { fst = refl ; snd = refl } ; otherFactorsNumber = 0 }
|
||||
|
||||
factorIntegerLemma (succ (succ x)) indHyp | a , record { fst = record { fst = divides record {quot = succ (succ qu) ; rem = .0 ; pr = ssxDivA ; remIsSmall = remSmall } refl ; snd = 1<a }; snd = record { fst = primeA ; snd = smallerFactors } } = inr record { 1<a = succPreservesInequality (succIsPositive x) ; firstFactor = a ; firstFactorNontrivial = Prime.p>1 primeA ; firstFactorBiggerMin = inl (Prime.p>1 primeA) ; firstFactorDivision = record { quot = succ (succ qu) ; rem = 0 ; pr = ssxDivA ; remIsSmall = remSmall ; quotSmall = inl (TotalOrder.transitive ℕTotalOrder (le zero refl) 1<a) } ; firstFactorDivides = refl ; firstFactorPrime = record { p>1 = Prime.p>1 primeA ; pr = Prime.pr primeA } ; otherFactors = inr record {fst = succPreservesInequality (succIsPositive qu) ; snd = factNonunit} ; otherFactorsNumber = succ (factorisationNonunit.otherFactorsNumber indHypRes') }
|
||||
@@ -143,7 +143,7 @@ module Numbers.Primes.IntegerFactorisation where
|
||||
... | indHyp = (indHyp prf) a factA p|a
|
||||
where
|
||||
prf : (x : ℕ) (ind : (y : ℕ) (y<x : y <N x) (factY : factorisationNonunit i y) (p|y : p ∣ y) → False) (factX : factorisationNonunit i x) (p|x : p ∣ x) → False
|
||||
prf x ind record { 1<a = 1<a ; firstFactor = firstFactor ; firstFactorNontrivial = firstFactorNontrivial ; firstFactorBiggerMin = firstFactorBiggerMin ; firstFactorDivision = firstFactorDivision ; firstFactorDivides = firstFactorDivides ; firstFactorPrime = firstFactorPrime ; otherFactors = (inl record { fst = quotFirstfact=1 ; snd = otherFactorsNumber }) } p|x = cannotBeLeqAndG i<=firstFactor p<i
|
||||
prf x ind record { 1<a = 1<a ; firstFactor = firstFactor ; firstFactorNontrivial = firstFactorNontrivial ; firstFactorBiggerMin = firstFactorBiggerMin ; firstFactorDivision = firstFactorDivision ; firstFactorDivides = firstFactorDivides ; firstFactorPrime = firstFactorPrime ; otherFactors = (inl record { fst = quotFirstfact=1 ; snd = otherFactorsNumber }) } p|x = exFalso bad
|
||||
where
|
||||
x=firstFact : firstFactor *N 1 +N 0 ≡ x
|
||||
x=firstFact rewrite equalityCommutative firstFactorDivides | equalityCommutative quotFirstfact=1 = divisionAlgResult.pr firstFactorDivision
|
||||
@@ -155,6 +155,10 @@ module Numbers.Primes.IntegerFactorisation where
|
||||
p=firstFact = primeDivPrimeImpliesEqual pPrime firstFactorPrime p|firstFact
|
||||
i<=firstFactor : i ≤N p
|
||||
i<=firstFactor rewrite p=firstFact = firstFactorBiggerMin
|
||||
bad : False
|
||||
bad with i<=firstFactor
|
||||
... | inl t = TotalOrder.irreflexive ℕTotalOrder (TotalOrder.transitive ℕTotalOrder t p<i)
|
||||
... | inr eq rewrite eq = TotalOrder.irreflexive ℕTotalOrder p<i
|
||||
prf zero ind record { 1<a = 1<a ; firstFactor = firstFactor ; firstFactorNontrivial = firstFactorNontrivial ; firstFactorBiggerMin = firstFactorBiggerMin ; firstFactorDivision = firstFactorDivision ; firstFactorDivides = firstFactorDivides ; firstFactorPrime = firstFactorPrime ; otherFactors = (inr otherFact) } p|x = zeroNeverGreater 1<a
|
||||
prf (succ x) ind record { 1<a = 1<a ; firstFactor = firstFactor ; firstFactorNontrivial = firstFactorNontrivial ; firstFactorBiggerMin = firstFactorBiggerMin ; firstFactorDivision = firstFactorDivision ; firstFactorDivides = firstFactorDivides ; firstFactorPrime = firstFactorPrime ; otherFactors = (inr otherFact) } p|x = ind (divisionAlgResult.quot firstFactorDivision) (lemma4 {divisionAlgResult.quot firstFactorDivision} {firstFactor} {divisionAlgResult.rem firstFactorDivision} {x} (divisionAlgResult.pr (firstFactorDivision)) (primesAreBiggerThanOne firstFactorPrime)) (canReduceFactorBound' (_&&_.snd otherFact) firstFactorBiggerMin) (p|q p|ffOrQ)
|
||||
where
|
||||
|
@@ -262,7 +262,7 @@ module Numbers.Primes.PrimeNumbers where
|
||||
biggerThanCantDivideLemma {zero} {b} a<b b|a = refl
|
||||
biggerThanCantDivideLemma {succ a} {zero} a<b (divides record { quot = quot ; rem = .0 ; pr = pr ; remIsSmall = remIsSmall ; quotSmall = (inl (le x ())) } refl)
|
||||
biggerThanCantDivideLemma {succ a} {zero} a<b (divides record { quot = quot ; rem = .0 ; pr = () ; remIsSmall = remIsSmall ; quotSmall = (inr x) } refl)
|
||||
biggerThanCantDivideLemma {succ a} {succ b} a<b (divides record { quot = zero ; rem = .0 ; pr = pr ; remIsSmall = remIsSmall ; quotSmall = quotSmall } refl) rewrite additionNIsCommutative (b *N zero) 0 | multiplicationNIsCommutative b 0 = naughtE pr
|
||||
biggerThanCantDivideLemma {succ a} {succ b} a<b (divides record { quot = zero ; rem = .0 ; pr = pr ; remIsSmall = remIsSmall ; quotSmall = quotSmall } refl) rewrite additionNIsCommutative (b *N zero) 0 | multiplicationNIsCommutative b 0 = exFalso (naughtE pr)
|
||||
biggerThanCantDivideLemma {succ a} {succ b} a<b (divides record { quot = (succ quot) ; rem = .0 ; pr = pr ; remIsSmall = remIsSmall ; quotSmall = quotSmall } refl) rewrite additionNIsCommutative (quot +N b *N succ quot) 0 | equalityCommutative pr = exFalso (positiveTimes {b} {quot} a<b)
|
||||
|
||||
biggerThanCantDivide : {a b : ℕ} → (x : ℕ) → (TotalOrder.max ℕTotalOrder a b) <N x → (x ∣ a) → (x ∣ b) → (a ≡ 0) && (b ≡ 0)
|
||||
@@ -346,7 +346,7 @@ module Numbers.Primes.PrimeNumbers where
|
||||
where
|
||||
lem : {b r : ℕ} → b *N r ≡ b → (0 <N b) → r ≡ 1
|
||||
lem {zero} {r} pr ()
|
||||
lem {succ b} {zero} pr _ rewrite multiplicationNIsCommutative b 0 = naughtE pr
|
||||
lem {succ b} {zero} pr _ rewrite multiplicationNIsCommutative b 0 = exFalso (naughtE pr)
|
||||
lem {succ b} {succ zero} pr _ = refl
|
||||
lem {succ b} {succ (succ r)} pr _ rewrite multiplicationNIsCommutative b (succ (succ r)) | additionNIsCommutative (succ r) (b +N (b +N r *N b)) | additionNIsAssociative b (b +N r *N b) (succ r) | additionNIsCommutative (b +N r *N b) (succ r) = exFalso (cannotAddAndEnlarge'' {succ b} pr)
|
||||
p : quot *N quotAB ≡ 1
|
||||
@@ -440,8 +440,7 @@ module Numbers.Primes.PrimeNumbers where
|
||||
euclidLemma2 {a} {b} {max} pr = lessTransitive {b} {succ (a +N b)} {max} (lemma a b) pr
|
||||
where
|
||||
lemma : (a b : ℕ) → b <N succ (a +N b)
|
||||
lemma a zero = succIsPositive (a +N zero)
|
||||
lemma a (succ b) = succPreservesInequality (collapseSuccOnRight {b} {a} {b} (lemma a b))
|
||||
lemma a b rewrite Semiring.commutative ℕSemiring (succ a) b = addingIncreases b a
|
||||
|
||||
euclidLemma3 : {a b max : ℕ} → (succ (succ (a +N b)) <N max) → succ b <N max
|
||||
euclidLemma3 {a} {b} {max} pr = euclidLemma2 {a} {succ b} {max} (identityOfIndiscernablesLeft (succ (succ (a +N b))) max (succ (a +N succ b)) _<N_ pr (applyEquality succ (equalityCommutative (succExtracts a b))))
|
||||
@@ -569,7 +568,7 @@ module Numbers.Primes.PrimeNumbers where
|
||||
f (divides record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = remIsSmall } x) rewrite x = naughtE pr
|
||||
divisionDecidable (succ a) b with divisionAlg (succ a) b
|
||||
divisionDecidable (succ a) b | record { quot = quot ; rem = zero ; pr = pr ; remIsSmall = remSmall } = inl (divides (record { quot = quot ; rem = zero ; pr = pr ; remIsSmall = remSmall ; quotSmall = inl (succIsPositive a) }) refl)
|
||||
divisionDecidable (succ a) b | record { quot = b/a ; rem = succ rem ; pr = prANotDivB ; remIsSmall = inr p } = naughtE (equalityCommutative p)
|
||||
divisionDecidable (succ a) b | record { quot = b/a ; rem = succ rem ; pr = prANotDivB ; remIsSmall = inr p } = exFalso (naughtE (equalityCommutative p))
|
||||
divisionDecidable (succ a) b | record { quot = b/a ; rem = succ rem ; pr = prANotDivB ; remIsSmall = inl p } = inr f
|
||||
where
|
||||
f : (succ a) ∣ b → False
|
||||
@@ -772,7 +771,7 @@ module Numbers.Primes.PrimeNumbers where
|
||||
q = succInjective (_&&_.snd p)
|
||||
|
||||
oneHasNoDivisors : {a : ℕ} → a ∣ 1 → a ≡ 1
|
||||
oneHasNoDivisors {a} (divides record { quot = zero ; rem = .0 ; pr = pr ; remIsSmall = remIsSmall } refl) rewrite addZeroRight (a *N zero) | multiplicationNIsCommutative a zero | addZeroRight a = naughtE pr
|
||||
oneHasNoDivisors {a} (divides record { quot = zero ; rem = .0 ; pr = pr ; remIsSmall = remIsSmall } refl) rewrite addZeroRight (a *N zero) | multiplicationNIsCommutative a zero | addZeroRight a = exFalso (naughtE pr)
|
||||
oneHasNoDivisors {a} (divides record { quot = (succ quot) ; rem = .0 ; pr = pr ; remIsSmall = remIsSmall } refl) rewrite addZeroRight (a *N succ quot) = _&&_.fst (mult1Lemma pr)
|
||||
|
||||
notSmallerMeansGE : {a b : ℕ} → (a <N b → False) → b ≤N a
|
||||
|
@@ -55,13 +55,13 @@ module Sets.CantorBijection.Proofs where
|
||||
cantorInverseOrderPreserving zero (succ y) x<y | inl (inl bl) = bl
|
||||
cantorInverseOrderPreserving zero (succ y) x<y | inl (inr bl) = exFalso (leastElement bl)
|
||||
cantorInverseOrderPreserving zero (succ y) x<y | inr x = exFalso (leastElement {cantorInverse y} (identityOfIndiscernablesRight _ _ _ order (cantorInverseLemmaIncreases (cantorInverse y)) (equalityCommutative x)))
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y with cantorInverseOrderPreserving x y (subtractOneFromInequality x<y)
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y with cantorInverseOrderPreserving x y (canRemoveSuccFrom<N x<y)
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr with cantorInverse x
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | x1 ,, x2 with cantorInverse y
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | zero ,, zero | zero ,, succ y2 = inl (succPreservesInequality (succIsPositive (y2 +N 0)))
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | zero ,, zero | succ y1 ,, zero with TotalOrder.totality ℕTotalOrder 1 (y1 +N 1)
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | zero ,, zero | succ y1 ,, zero | inl (inl pr1) = inl pr1
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | zero ,, zero | succ y1 ,, zero | inl (inr pr1) rewrite Semiring.commutative ℕSemiring y1 1 = exFalso (zeroNeverGreater (subtractOneFromInequality pr1))
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | zero ,, zero | succ y1 ,, zero | inl (inr pr1) rewrite Semiring.commutative ℕSemiring y1 1 = exFalso (zeroNeverGreater (canRemoveSuccFrom<N pr1))
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | zero ,, zero | succ y1 ,, zero | inr pr1 = inr (pr1 ,, le 0 refl)
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | zero ,, zero | succ y1 ,, succ y2 rewrite Semiring.commutative ℕSemiring y1 (succ (succ y2)) = inl (succPreservesInequality (succIsPositive (y2 +N y1)))
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | succ x1 ,, zero | zero ,, succ y2 rewrite Semiring.commutative ℕSemiring x1 1 | Semiring.sumZeroRight ℕSemiring y2 | Semiring.sumZeroRight ℕSemiring x1 = inl (TotalOrder.transitive ℕTotalOrder pr (a<SuccA _))
|
||||
@@ -106,23 +106,23 @@ module Sets.CantorBijection.Proofs where
|
||||
cantorInverseDiscrete : (a : ℕ) → (c : ℕ && ℕ) → order (cantorInverse a) c → order c (cantorInverse (succ a)) → False
|
||||
cantorInverseDiscrete zero (zero ,, zero) (inl ()) c<sa
|
||||
cantorInverseDiscrete zero (zero ,, zero) (inr ()) c<sa
|
||||
cantorInverseDiscrete zero (zero ,, succ c) a<c (inl x) = zeroNeverGreater (subtractOneFromInequality x)
|
||||
cantorInverseDiscrete zero (succ b ,, zero) a<c (inl x) = zeroNeverGreater (subtractOneFromInequality x)
|
||||
cantorInverseDiscrete zero (succ b ,, succ c) a<c (inl x) = zeroNeverGreater (subtractOneFromInequality x)
|
||||
cantorInverseDiscrete zero (zero ,, succ c) a<c (inl x) = zeroNeverGreater (canRemoveSuccFrom<N x)
|
||||
cantorInverseDiscrete zero (succ b ,, zero) a<c (inl x) = zeroNeverGreater (canRemoveSuccFrom<N x)
|
||||
cantorInverseDiscrete zero (succ b ,, succ c) a<c (inl x) = zeroNeverGreater (canRemoveSuccFrom<N x)
|
||||
cantorInverseDiscrete (succ a) (b ,, c) a<c c<sa with cantorInverse a
|
||||
cantorInverseDiscrete (succ a) (zero ,, zero) (inl ()) (inl x) | zero ,, zero
|
||||
cantorInverseDiscrete (succ a) (zero ,, zero) (inr ()) (inl x) | zero ,, zero
|
||||
cantorInverseDiscrete (succ a) (zero ,, succ zero) a<c (inl x) | zero ,, zero = TotalOrder.irreflexive ℕTotalOrder x
|
||||
cantorInverseDiscrete (succ a) (zero ,, succ (succ c)) a<c (inl x) | zero ,, zero = zeroNeverGreater (subtractOneFromInequality x)
|
||||
cantorInverseDiscrete (succ a) (succ b ,, c) a<c (inl x) | zero ,, zero = zeroNeverGreater (subtractOneFromInequality x)
|
||||
cantorInverseDiscrete (succ a) (zero ,, succ (succ c)) a<c (inl x) | zero ,, zero = zeroNeverGreater (canRemoveSuccFrom<N x)
|
||||
cantorInverseDiscrete (succ a) (succ b ,, c) a<c (inl x) | zero ,, zero = zeroNeverGreater (canRemoveSuccFrom<N x)
|
||||
cantorInverseDiscrete (succ a) (b ,, zero) (inl x) (inr (fst ,, snd)) | zero ,, zero = TotalOrder.irreflexive ℕTotalOrder (identityOfIndiscernablesRight _ _ _ _<N_ x fst)
|
||||
cantorInverseDiscrete (succ a) (b ,, succ c) a<c (inr (fst ,, snd)) | zero ,, zero = zeroNeverGreater (subtractOneFromInequality snd)
|
||||
cantorInverseDiscrete (succ a) (b ,, succ c) a<c (inr (fst ,, snd)) | zero ,, zero = zeroNeverGreater (canRemoveSuccFrom<N snd)
|
||||
cantorInverseDiscrete (succ a) (b ,, c) (inl y) (inl x) | zero ,, succ snd rewrite Semiring.commutative ℕSemiring snd 1 | Semiring.sumZeroRight ℕSemiring snd = TotalOrder.irreflexive ℕTotalOrder (TotalOrder.transitive ℕTotalOrder x y)
|
||||
cantorInverseDiscrete (succ a) (b ,, succ c) (inr (fst ,, _)) (inl x) | zero ,, succ snd rewrite Semiring.commutative ℕSemiring snd 1 | Semiring.sumZeroRight ℕSemiring snd | fst = TotalOrder.irreflexive ℕTotalOrder x
|
||||
cantorInverseDiscrete (succ a) (b ,, zero) (inl x) (inr (fst ,, snd₁)) | zero ,, succ snd rewrite fst | Semiring.commutative ℕSemiring snd 1 | Semiring.sumZeroRight ℕSemiring snd = TotalOrder.irreflexive ℕTotalOrder x
|
||||
cantorInverseDiscrete (succ a) (b ,, succ c) (inl x) (inr (fst ,, snd1)) | zero ,, succ snd = zeroNeverGreater (subtractOneFromInequality snd1)
|
||||
cantorInverseDiscrete (succ a) (b ,, succ c) (inl x) (inr (fst ,, snd1)) | zero ,, succ snd = zeroNeverGreater (canRemoveSuccFrom<N snd1)
|
||||
cantorInverseDiscrete (succ a) (b ,, succ zero) (inr (fst ,, _)) (inr (fst₁ ,, bad)) | zero ,, succ snd = TotalOrder.irreflexive ℕTotalOrder bad
|
||||
cantorInverseDiscrete (succ a) (b ,, succ (succ c)) (inr (fst ,, _)) (inr (fst₁ ,, bad)) | zero ,, succ snd = zeroNeverGreater (subtractOneFromInequality bad)
|
||||
cantorInverseDiscrete (succ a) (b ,, succ (succ c)) (inr (fst ,, _)) (inr (fst₁ ,, bad)) | zero ,, succ snd = zeroNeverGreater (canRemoveSuccFrom<N bad)
|
||||
cantorInverseDiscrete (succ a) (b ,, c) (inl x1) (inl x) | succ zero ,, zero = noIntegersBetweenXAndSuccX 1 x1 x
|
||||
cantorInverseDiscrete (succ a) (zero ,, succ zero) (inr (fst ,, snd)) (inl x) | succ zero ,, zero = TotalOrder.irreflexive ℕTotalOrder snd
|
||||
cantorInverseDiscrete (succ a) (succ b ,, succ c) (inr (fst ,, snd)) (inl x) | succ zero ,, zero rewrite Semiring.commutative ℕSemiring b (succ c) = bad fst
|
||||
|
Reference in New Issue
Block a user