First Git commit

This commit is contained in:
Smaug123
2019-01-04 20:45:34 +00:00
commit a435e3764e
39 changed files with 9074 additions and 0 deletions

231
Cardinality.agda Normal file
View File

@@ -0,0 +1,231 @@
{-# OPTIONS --safe --warning=error #-}
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import LogicalFormulae
open import Functions
open import Naturals
open import FinSet
module Cardinality where
record CountablyInfiniteSet {a : _} (A : Set a) : Set a where
field
counting : A
countingIsBij : Bijection counting
data Countable {a : _} (A : Set a) : Set a where
finite : FiniteSet A Countable A
infinite : CountablyInfiniteSet A Countable A
Countable : CountablyInfiniteSet
Countable = record { counting = id ; countingIsBij = invertibleImpliesBijection (record { inverse = id ; isLeft = λ b refl ; isRight = λ a refl}) }
doubleLemma : (a b : ) 2 *N a 2 *N b a b
doubleLemma a b pr = productCancelsLeft 2 a b (le 1 refl) pr
evenCannotBeOdd : (a b : ) 2 *N a succ (2 *N b) False
evenCannotBeOdd zero b ()
evenCannotBeOdd (succ a) zero pr rewrite additionNIsCommutative a 0 | additionNIsCommutative a (succ a) = naughtE (equalityCommutative (succInjective pr))
evenCannotBeOdd (succ a) (succ b) pr = evenCannotBeOdd a b pr''
where
pr' : a +N a (b +N succ b)
pr' rewrite additionNIsCommutative a 0 | additionNIsCommutative b 0 | additionNIsCommutative a (succ a) = succInjective (succInjective pr)
pr'' : 2 *N a succ (2 *N b)
pr'' rewrite additionNIsCommutative a 0 | additionNIsCommutative b 0 | additionNIsCommutative (succ b) b = pr'
aMod2 : (a : ) Sg (λ i (2 *N i a) || (succ (2 *N i) a))
aMod2 zero = (0 , inl refl)
aMod2 (succ a) with aMod2 a
aMod2 (succ a) | b , inl x = b , inr (applyEquality succ x)
aMod2 (succ a) | b , inr x = (succ b) , inl pr
where
pr : succ (b +N succ (b +N 0)) succ a
pr rewrite additionNIsCommutative b (succ (b +N 0)) | additionNIsCommutative (b +N 0) b = applyEquality succ x
sqrtFloor : (a : ) Sg ( && ) (λ pair ((_&&_.fst pair) *N (_&&_.fst pair) +N (_&&_.snd pair) a) && ((_&&_.snd pair) <N 2 *N (_&&_.fst pair) +N 1))
sqrtFloor zero = (0 ,, 0) , (refl ,, le zero refl)
sqrtFloor (succ n) with sqrtFloor n
sqrtFloor (succ n) | (a ,, b) , pr with orderIsTotal b (2 *N a)
sqrtFloor (succ n) | (a ,, b) , pr | inl (inl x) = (a ,, succ b) , (p ,, q)
where
p : a *N a +N succ b succ n
p rewrite additionNIsCommutative (a *N a) (succ b) | additionNIsCommutative b (a *N a) = applyEquality succ (_&&_.fst pr)
q : succ b <N (a +N (a +N 0)) +N 1
q rewrite additionNIsCommutative (a +N (a +N 0)) (succ 0) | additionNIsCommutative a 0 = succPreservesInequality x
sqrtFloor (succ n) | (a ,, b) , (_ ,, pr) | inl (inr x) rewrite additionNIsCommutative (a +N (a +N 0)) (succ 0) = exFalso (noIntegersBetweenXAndSuccX (a +N (a +N 0)) x pr)
sqrtFloor (succ n) | (a ,, b) , pr | inr x = (succ a ,, 0) , (q ,, succIsPositive _)
where
p : a +N a *N succ a n
p rewrite x | additionNIsCommutative a 0 | additionNIsCommutative (a +N a) (succ 0) | additionNIsCommutative (a *N a) (succ a +N a) | multiplicationNIsCommutative a (succ a) | additionNIsCommutative (a *N a) (a +N a) | additionNIsAssociative a a (a *N a) = _&&_.fst pr
q : succ ((a +N a *N succ a) +N 0) succ n
q rewrite additionNIsCommutative (a +N a *N succ a) 0 = applyEquality succ p
{-
triangle : (a : ) →
triangle zero = 0
triangle (succ a) = succ a +N triangle a
cantorPairing' : (a : ) → (b : ) → (s : ) → (s ≡ a +N b) →
cantorPairing' a b zero pr = 0
cantorPairing' zero zero (succ s) ()
cantorPairing' zero (succ b) (succ s) pr = succ (cantorPairing' 1 b (succ s) pr)
cantorPairing' (succ a) zero (succ s) pr = succ (cantorPairing' 0 a s (succInjective (identityOfIndiscernablesRight _ _ _ _≡_ pr (applyEquality (λ i → succ i) (identityOfIndiscernablesLeft _ _ _ _≡_ (additionNIsCommutative a 0) refl)))))
cantorPairing' (succ a) (succ b) (succ zero) pr = naughtE f
where
f : 0 ≡ succ b +N a
f rewrite additionNIsCommutative (succ b) a = succInjective pr
cantorPairing' (succ a) (succ b) (succ (succ s)) pr = succ (cantorPairing' (succ (succ a)) b (succ (succ s)) (identityOfIndiscernablesRight _ _ _ _≡_ pr (applyEquality succ (succExtracts a b))))
cantorPairing'Zero : (x y s : ) → (pr : s ≡ x +N y) → cantorPairing' x y s pr ≡ 0 → (x ≡ 0) && (y ≡ 0)
cantorPairing'Zero x y zero pr pr' = sumZeroImpliesSummandsZero (equalityCommutative pr)
cantorPairing'Zero zero zero (succ s) () pr'
cantorPairing'Zero zero (succ y) (succ s) pr ()
cantorPairing'Zero (succ x) zero (succ s) pr ()
cantorPairing'Zero (succ x) (succ y) (succ zero) pr pr' = naughtE (identityOfIndiscernablesRight _ _ _ _≡_ (succInjective pr) (additionNIsCommutative x (succ y)))
cantorPairing'Zero (succ x) (succ y) (succ (succ s)) pr ()
cantorPairingInj' : (s1 s2 : ) → (x1 x2 y1 y2 : ) → (pr1 : s1 ≡ x1 +N y1) → (pr2 : s2 ≡ x2 +N y2) → (cantorPairing' x1 y1 s1 pr1) ≡ (cantorPairing' x2 y2 s2 pr2) → (x1 ≡ x2) && (y1 ≡ y2)
cantorPairingInj' zero zero x1 x2 y1 y2 pr1 pr2 injF = transitivity (_&&_.fst x1y1) (equalityCommutative (_&&_.fst x2y2)) ,, transitivity (_&&_.snd x1y1) (equalityCommutative (_&&_.snd x2y2))
where
x2y2 : (x2 ≡ 0) && (y2 ≡ 0)
x2y2 = sumZeroImpliesSummandsZero (equalityCommutative pr2)
x1y1 : (x1 ≡ 0) && (y1 ≡ 0)
x1y1 = sumZeroImpliesSummandsZero (equalityCommutative pr1)
cantorPairingInj' zero (succ s2) x1 x2 y1 y2 pr1 pr2 injF = naughtE (equalityCommutative bad2)
where
x2y2 : (x2 ≡ 0) && (y2 ≡ 0)
x2y2 = cantorPairing'Zero x2 y2 (succ s2) pr2 (equalityCommutative injF)
bad : succ s2 ≡ y2
bad = identityOfIndiscernablesRight _ _ _ _≡_ pr2 ((applyEquality (λ i → i +N y2) (_&&_.fst x2y2)))
bad2 : succ s2 ≡ 0
bad2 = transitivity bad (_&&_.snd x2y2)
cantorPairingInj' (succ s1) zero x1 x2 y1 y2 pr1 pr2 injF = naughtE (equalityCommutative bad2)
where
x1y1 : (x1 ≡ 0) && (y1 ≡ 0)
x1y1 = cantorPairing'Zero x1 y1 (succ s1) pr1 injF
bad : succ s1 ≡ y1
bad = identityOfIndiscernablesRight _ _ _ _≡_ pr1 ((applyEquality (λ i → i +N y1) (_&&_.fst x1y1)))
bad2 : succ s1 ≡ 0
bad2 = transitivity bad (_&&_.snd x1y1)
cantorPairingInj' (succ s1) (succ s2) zero x2 zero y2 () pr2 injF
cantorPairingInj' (succ s1) (succ s2) zero zero (succ .s1) zero refl () injF
cantorPairingInj' (succ zero) (succ zero) zero zero (succ .0) (succ .0) refl refl injF = refl ,, refl
cantorPairingInj' (succ zero) (succ (succ s2)) zero zero (succ .0) (succ .(succ s2)) refl refl injF = naughtE (equalityCommutative bad)
where
bad : 2 ≡ 0
bad = _&&_.fst (cantorPairing'Zero 2 s2 (succ (succ s2)) refl (equalityCommutative (succInjective (succInjective injF))))
cantorPairingInj' (succ (succ s1)) (succ zero) zero zero (succ .(succ s1)) (succ .0) refl refl injF = naughtE (equalityCommutative bad)
where
bad : 2 ≡ 0
bad = _&&_.fst (cantorPairing'Zero 2 s1 (succ (succ s1)) refl (succInjective (succInjective injF)))
cantorPairingInj' (succ (succ s1)) (succ (succ s2)) zero zero (succ .(succ s1)) (succ .(succ s2)) refl refl injF = refl ,, (applyEquality (λ i → succ (succ i)) rec')
where
rec : cantorPairing' 2 s1 (succ (succ s1)) refl ≡ cantorPairing' 2 s2 (succ (succ s2)) refl
rec = succInjective (succInjective injF)
rec' : (s1 ≡ s2)
rec' = _&&_.snd (cantorPairingInj' (succ (succ s1)) (succ (succ s2)) 2 2 s1 s2 refl refl rec)
cantorPairingInj' (succ s1) (succ s2) zero (succ x2) (succ .s1) zero refl pr2 injF = naughtE (equalityCommutative rec')
where
rec : cantorPairing' 1 s1 (succ s1) refl ≡ cantorPairing' 0 x2 s2 _
rec = succInjective injF
rec' : (1 ≡ zero)
rec' = _&&_.fst (cantorPairingInj' (succ s1) s2 1 zero s1 x2 refl _ rec)
cantorPairingInj' (succ s1) (succ zero) zero (succ x2) (succ .s1) (succ y2) refl pr2 injF = naughtE (identityOfIndiscernablesRight _ _ _ _≡_ (succInjective pr2) (additionNIsCommutative x2 (succ y2)))
cantorPairingInj' (succ s1) (succ (succ s2)) zero (succ x2) (succ .s1) (succ y2) refl pr2 injF = naughtE (succInjective rec')
where
rec : cantorPairing' 1 s1 (succ s1) refl ≡ cantorPairing' (succ (succ x2)) y2 (succ (succ s2)) _
rec = succInjective injF
rec' : (1 ≡ succ (succ x2))
rec' = _&&_.fst (cantorPairingInj' (succ s1) (succ (succ s2)) 1 (succ (succ x2)) s1 y2 refl _ rec)
cantorPairingInj' (succ s1) (succ s2) (succ x1) zero y1 zero pr1 () injF
cantorPairingInj' (succ s1) (succ s2) (succ zero) zero .s1 (succ .s2) refl refl injF = {!!}
cantorPairingInj' (succ s1) (succ s2) (succ (succ x1)) zero zero (succ .s2) pr1 refl injF = {!!}
where
false : {!!}
false = {!!}
cantorPairingInj' (succ s1) (succ s2) (succ (succ x1)) zero (succ y1) (succ .s2) pr1 refl injF = {!!}
cantorPairingInj' (succ s1) (succ s2) (succ x1) (succ x2) y1 y2 pr1 pr2 injF = {!!}
cantorPairingInj : Injection cantorPairing
Injection.property cantorPairingInj {zero ,, zero} {zero ,, zero} pr = refl
Injection.property cantorPairingInj {zero ,, zero} {zero ,, succ snd} pr = exFalso {!!}
Injection.property cantorPairingInj {zero ,, zero} {succ fst ,, snd} pr = {!!}
Injection.property cantorPairingInj {zero ,, succ snd} {y} pr = {!!}
Injection.property cantorPairingInj {succ fst ,, snd} {y} pr = {!!}
cantorInverse : → ( && )
cantorInverse zero = (0 ,, 0)
cantorInverse (succ z) = f (cantorInverse z)
where
f : ( && ) → ( && )
f (0 ,, s) = (succ s) ,, 0
f (succ r ,, 0) = r ,, 1
f (succ r ,, succ s) = (r ,, succ (succ s))
cantorInvertible : Invertible cantorPairing
Invertible.inverse cantorInvertible = cantorInverse
Invertible.isLeft cantorInvertible zero = refl
Invertible.isLeft cantorInvertible (succ b) with inspect (cantorInverse b)
Invertible.isLeft cantorInvertible (succ b) | (zero ,, zero) with≡ x = {!!}
Invertible.isLeft cantorInvertible (succ b) | (zero ,, succ snd) with≡ x = {!!}
Invertible.isLeft cantorInvertible (succ b) | (succ fst ,, snd) with≡ x = {!!}
Invertible.isRight cantorInvertible (a ,, b) = {!!}
cantorBijection : Bijection cantorPairing
cantorBijection = invertibleImpliesBijection cantorInvertible
-}
pairUnionIsCountable : {a b : _} {A : Set a} {B : Set b} (X : CountablyInfiniteSet A) (Y : CountablyInfiniteSet B) CountablyInfiniteSet (A || B)
CountablyInfiniteSet.counting (pairUnionIsCountable X Y) (inl r) = 2 *N (CountablyInfiniteSet.counting X r)
CountablyInfiniteSet.counting (pairUnionIsCountable X Y) (inr s) = succ (2 *N (CountablyInfiniteSet.counting Y s))
Injection.property (Bijection.inj (CountablyInfiniteSet.countingIsBij (pairUnionIsCountable X Y))) {inl x} {inl y} pr rewrite additionNIsCommutative (CountablyInfiniteSet.counting X x) 0 | additionNIsCommutative (CountablyInfiniteSet.counting X y) 0 | doubleIsAddTwo (CountablyInfiniteSet.counting X x) | doubleIsAddTwo (CountablyInfiniteSet.counting X y) = applyEquality inl (Injection.property (Bijection.inj (CountablyInfiniteSet.countingIsBij X)) inter)
where
inter : CountablyInfiniteSet.counting X x CountablyInfiniteSet.counting X y
inter = doubleLemma (CountablyInfiniteSet.counting X x) (CountablyInfiniteSet.counting X y) pr
Injection.property (Bijection.inj (CountablyInfiniteSet.countingIsBij (pairUnionIsCountable X Y))) {inl x} {inr y} pr = exFalso (evenCannotBeOdd (CountablyInfiniteSet.counting X x) (CountablyInfiniteSet.counting Y y) pr)
Injection.property (Bijection.inj (CountablyInfiniteSet.countingIsBij (pairUnionIsCountable X Y))) {inr x} {inl y} pr = exFalso (evenCannotBeOdd (CountablyInfiniteSet.counting X y) (CountablyInfiniteSet.counting Y x) (equalityCommutative pr))
Injection.property (Bijection.inj (CountablyInfiniteSet.countingIsBij (pairUnionIsCountable X Y))) {inr x} {inr y} pr = applyEquality inr (Injection.property (Bijection.inj (CountablyInfiniteSet.countingIsBij Y)) (doubleLemma (CountablyInfiniteSet.counting Y x) (CountablyInfiniteSet.counting Y y) (succInjective pr) ))
Surjection.property (Bijection.surj (CountablyInfiniteSet.countingIsBij (pairUnionIsCountable X Y))) b with aMod2 b
Surjection.property (Bijection.surj (CountablyInfiniteSet.countingIsBij (pairUnionIsCountable X Y))) b | a , inl x with Surjection.property (Bijection.surj (CountablyInfiniteSet.countingIsBij X)) a
Surjection.property (Bijection.surj (CountablyInfiniteSet.countingIsBij (pairUnionIsCountable X Y))) b | a , inl x | r , pr = inl r , ans
where
ans : 2 *N CountablyInfiniteSet.counting X r b
ans rewrite pr = x
Surjection.property (Bijection.surj (CountablyInfiniteSet.countingIsBij (pairUnionIsCountable X Y))) b | a , inr x with Surjection.property (Bijection.surj (CountablyInfiniteSet.countingIsBij Y)) a
Surjection.property (Bijection.surj (CountablyInfiniteSet.countingIsBij (pairUnionIsCountable X Y))) b | a , inr x | r , pr = inr r , ans
where
ans : succ (2 *N CountablyInfiniteSet.counting Y r) b
ans rewrite pr = x
firstEqualityOfPair : {a b : _} {A : Set a} {B : Set b} {x1 x2 : A} {y1 y2 : B} (x1 ,, y1) (x2 ,, y2) x1 x2
firstEqualityOfPair {x1} {x2} {y1} {y2} refl = refl
secondEqualityOfPair : {a b : _} {A : Set a} {B : Set b} {x1 x2 : A} {y1 y2 : B} (x1 ,, y1) (x2 ,, y2) y1 y2
secondEqualityOfPair {x1} {x2} {y1} {y2} refl = refl
{-
pairProductIsCountable : {a b : _} {A : Set a} {B : Set b} → (X : CountablyInfiniteSet A) → (Y : CountablyInfiniteSet B) → CountablyInfiniteSet (A && B)
CountablyInfiniteSet.counting (pairProductIsCountable X Y) (a ,, b) = cantorPairing record { fst = CountablyInfiniteSet.counting X a ; snd = CountablyInfiniteSet.counting Y b }
Injection.property (Bijection.inj (CountablyInfiniteSet.countingIsBij (pairProductIsCountable {a} {b} {A} {B} X Y))) {(x1 ,, y1)} {(x2 ,, y2)} pr = ans x1=x2 y1=y2
where
p : (CountablyInfiniteSet.counting X x1 ,, CountablyInfiniteSet.counting Y y1) ≡ (CountablyInfiniteSet.counting X x2 ,, CountablyInfiniteSet.counting Y y2)
p = Injection.property (Bijection.inj cantorBijection) pr
p1 : CountablyInfiniteSet.counting X x1 ≡ CountablyInfiniteSet.counting X x2
p1 = firstEqualityOfPair p
p2 : CountablyInfiniteSet.counting Y y1 ≡ CountablyInfiniteSet.counting Y y2
p2 = secondEqualityOfPair p
x1=x2 : x1 ≡ x2
x1=x2 = Injection.property (Bijection.inj (CountablyInfiniteSet.countingIsBij X)) p1
y1=y2 : y1 ≡ y2
y1=y2 = Injection.property (Bijection.inj (CountablyInfiniteSet.countingIsBij Y)) p2
ans : {a b : _} {A : Set a} {B : Set b} {x1 x2 : A} {y1 y2 : B} → (x1 ≡ x2) → (y1 ≡ y2) → (x1 ,, y1) ≡ (x2 ,, y2)
ans x1=x2 y1=y2 rewrite x1=x2 | y1=y2 = refl
Surjection.property (Bijection.surj (CountablyInfiniteSet.countingIsBij (pairProductIsCountable X Y))) b with inspect (cantorInverse b)
Surjection.property (Bijection.surj (CountablyInfiniteSet.countingIsBij (pairProductIsCountable X Y))) b | (n1 ,, n2) with≡ blah with Surjection.property (Bijection.surj (CountablyInfiniteSet.countingIsBij X)) n1
Surjection.property (Bijection.surj (CountablyInfiniteSet.countingIsBij (pairProductIsCountable X Y))) b | (n1 ,, n2) with≡ blah | b1 , pr1 with Surjection.property (Bijection.surj (CountablyInfiniteSet.countingIsBij Y)) n2
... | b2 , pr2 = (b1 ,, b2) , {!!}
-}
injectionToNMeansCountable : {a : _} {A : Set a} {f : A } Injection f InfiniteSet A Countable A
injectionToNMeansCountable {f = f} inj record { isInfinite = isInfinite } = {!!}

109
Category.agda Normal file
View File

@@ -0,0 +1,109 @@
{-# OPTIONS --warning=error #-}
open import LogicalFormulae
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Naturals
open import Vectors
postulate
extensionality : {a b : _} {S : Set a}{T : S Set b} {f g : (x : S) T x} ((x : S) f x g x) f g
record Category {a b : _} : Set (lsuc (a b)) where
field
objects : Set a
arrows : objects objects Set b
id : (x : objects) arrows x x
_∘_ : {x y z : objects} arrows y z arrows x y arrows x z
rightId : {x y : objects} (f : arrows x y) (id y) f f
leftId : {x y : objects} (f : arrows x y) f (id x) f
associative : {x y z w : objects} (f : arrows z w) (g : arrows y z) (h : arrows x y) (f (g h)) (f g) h
dual : {a b : _} Category {a} {b} Category {a} {b}
dual record { objects = objects ; arrows = arrows ; id = id ; _∘_ = _∘_ ; rightId = rightId ; leftId = leftId ; associative = associative } = record { objects = objects ; arrows = λ i j arrows j i ; id = id ; _∘_ = λ {x y z} g f f g ; rightId = λ {x y} f leftId f ; leftId = λ {x y} f rightId f ; associative = λ {x y z w} f g h equalityCommutative (associative h g f) }
SET : {a : _} Category {lsuc a} {a}
SET {a} = record { objects = Set a ; arrows = λ i j (i j) ; id = λ X x x ; _∘_ = λ g f x g (f x) ; rightId = λ f refl ; leftId = λ f refl ; associative = λ f g h refl }
≡Unique : {a : _} {X : Set a} {a b : X} (p1 p2 : a b) (p1 p2)
≡Unique refl refl = refl
NatPreorder : Category {lzero} {lzero}
NatPreorder = record { objects = ; arrows = λ m n m ≤N n ; id = λ x inr refl ; _∘_ = λ f g leqTransitive g f ; rightId = λ x<y leqUnique (leqTransitive x<y (inr refl)) x<y ; leftId = λ x<y leqUnique (leqTransitive (inr refl) x<y) x<y ; associative = λ z<=w y<=z x<=y leqUnique (leqTransitive (leqTransitive x<=y y<=z) z<=w) (leqTransitive x<=y (leqTransitive y<=z z<=w)) }
where
leqTransitive : {a b c : } (a ≤N b) (b ≤N c) (a ≤N c)
leqTransitive (inl a<b) (inl b<c) = inl (orderIsTransitive a<b b<c)
leqTransitive (inl a<b) (inr b=c) rewrite b=c = inl a<b
leqTransitive (inr a=b) (inl b<c) rewrite a=b = inl b<c
leqTransitive (inr a=b) (inr b=c) rewrite a=b | b=c = inr refl
<Nunique : {a b : } (p1 p2 : a <N b) p1 p2
<Nunique {a} {b} (le a-b pr1) (le a-b2 pr2) = go a-b pr1 a-b2 pr2 p'
where
p : a-b2 +N a a-b +N a
p rewrite equalityCommutative pr1 = succInjective pr2
p' : a-b2 a-b
p' = canSubtractFromEqualityRight p
go : (a-b : ) (pr1 : succ (a-b +N a) b) (a-b2 : ) (pr2 : succ (a-b2 +N a) b) (p : a-b2 a-b) (le a-b pr1) (le a-b2 pr2)
go a-b pr1 a-b2 pr2 eq rewrite eq = applyEquality (λ i le a-b i) (≡Unique pr1 pr2)
leqUnique : {a b : } (p1 : a ≤N b) (p2 : a ≤N b) p1 p2
leqUnique (inl a<b) (inl a<b2) = applyEquality inl (<Nunique a<b a<b2)
leqUnique (inl a<b) (inr a=b) rewrite a=b = exFalso (lessIrreflexive a<b)
leqUnique (inr a=b) (inl a<b) rewrite a=b = exFalso (lessIrreflexive a<b)
leqUnique (inr a=b1) (inr a=b2) rewrite a=b1 | a=b2 = refl
NatMonoid : Category {lzero} {lzero}
NatMonoid = record { objects = True ; arrows = λ _ _ ; id = λ x 0 ; _∘_ = λ f g f +N g ; rightId = λ f refl ; leftId = λ f addZeroRight f ; associative = λ a b c equalityCommutative (additionNIsAssociative a b c) }
DISCRETE : {a : _} (X : Set a) Category {a} {a}
DISCRETE X = record { objects = X ; arrows = _≡_ ; id = λ x refl ; _∘_ = λ f g transitivity g f ; rightId = λ f ≡Unique (transitivity f refl) f ; leftId = λ f ≡Unique (transitivity refl f) f ; associative = λ f g h ≡Unique (transitivity (transitivity h g) f) (transitivity h (transitivity g f)) }
record Functor {a b c d : _} (C : Category {a} {b}) (D : Category {c} {d}) : Set (a b c d) where
field
onObj : Category.objects C Category.objects D
onArrow : {S T : Category.objects C} Category.arrows C S T Category.arrows D (onObj S) (onObj T)
mapId : {T : Category.objects C} onArrow (Category.id C T) Category.id D (onObj T)
mapCompose : {X Y Z : Category.objects C} (f : Category.arrows C X Y) (g : Category.arrows C Y Z) onArrow (Category._∘_ C g f) Category._∘_ D (onArrow g) (onArrow f)
functorCompose : {a b c d e f : _} {B : Category {a} {b}} {C : Category {c} {d}} {D : Category {e} {f}} (Functor C D) (Functor B C) (Functor B D)
functorCompose G F = record { onObj = λ x Functor.onObj G (Functor.onObj F x) ; onArrow = λ f Functor.onArrow G (Functor.onArrow F f) ; mapId = λ {T} mapIdHelp G F T ; mapCompose = λ r s mapComposeHelp G F r s }
where
mapIdHelp : {a b c d e f : _} {B : Category {a} {b}} {C : Category {c} {d}} {D : Category {e} {f}} (G : Functor C D) (F : Functor B C) (T : Category.objects B) Functor.onArrow G (Functor.onArrow F (Category.id B T)) Category.id D (Functor.onObj G (Functor.onObj F T))
mapIdHelp {B = B} {C} {D} record { onObj = onObjG ; onArrow = onArrowG ; mapId = mapIdG ; mapCompose = mapComposeG } record { onObj = onObj ; onArrow = onArrow ; mapId = mapId ; mapCompose = mapCompose } T rewrite mapId {T} = mapIdG {onObj T}
mapComposeHelp : {a b c d e f : _} {B : Category {a} {b}} {C : Category {c} {d}} {D : Category {e} {f}} (G : Functor C D) (F : Functor B C) {S T U : Category.objects B} (r : Category.arrows B S T) (s : Category.arrows B T U) (Functor.onArrow G (Functor.onArrow F (Category._∘_ B s r))) (Category._∘_ D (Functor.onArrow G (Functor.onArrow F s)) (Functor.onArrow G (Functor.onArrow F r)))
mapComposeHelp {B = record { objects = objectsB ; arrows = arrowsB ; id = idB ; _∘_ = _∘B_ ; rightId = rightIdB ; leftId = leftIdB ; associative = associativeB }} {record { objects = objectsC ; arrows = arrowsC ; id = idC ; _∘_ = _∘C_ ; rightId = rightIdC ; leftId = leftIdC ; associative = associativeC }} {record { objects = objectsD ; arrows = arrowsD ; id = idD ; _∘_ = _∘D_ ; rightId = rightIdD ; leftId = leftIdD ; associative = associativeD }} record { onObj = onObjG ; onArrow = onArrowG ; mapId = mapIdG ; mapCompose = mapComposeG } record { onObj = onObjF ; onArrow = onArrowF ; mapId = mapIdF ; mapCompose = mapComposeF } {S} {T} {U} r s rewrite mapComposeF r s | mapComposeG (onArrowF r) (onArrowF s) = refl
idFunctor : {a b : _} (C : Category {a} {b}) Functor C C
Functor.onObj (idFunctor C) = λ x x
Functor.onArrow (idFunctor C) = λ f f
Functor.mapId (idFunctor C) = refl
Functor.mapCompose (idFunctor C) = λ f g refl
typeCastCat : {a b c d : _} {C : Category {a} {b}} {D : Category {c} {d}} (F : Functor C D) (G : Functor C D) (S T : Category.objects C) (pr : Functor.onObj F Functor.onObj G) Category.arrows D (Functor.onObj G S) (Functor.onObj G T) Category.arrows D (Functor.onObj F S) (Functor.onObj F T)
typeCastCat F G S T pr rewrite pr = refl
equalityFunctionsEqual : {a b : _} {A : Set a} {B : Set b} (f : A (B B)) (g : A (B B)) (f g)
equalityFunctionsEqual f g = extensionality λ x ≡Unique (f x) (g x)
equalityFunctionsEqual' : {a b : _} {A : Set a} {B : Set b} (f : A (B B)) (g : A (B B)) (f g)
equalityFunctionsEqual' f g = extensionality λ x ≡Unique (f x) (g x)
functorsEqual' : {a b c d : _} {C : Category {a} {b}} {D : Category {c} {d}} (F : Functor C D) (G : Functor C D) (objEq : (Functor.onObj F) Functor.onObj G) (arrEq : {S T : Category.objects C} {f : Category.arrows C S T} (Functor.onArrow F {S} {T} f (typeCast (Functor.onArrow G {S} {T} f) (typeCastCat F G S T objEq)))) F G
functorsEqual' record { onObj = onObjF ; onArrow = onArrowF ; mapId = mapIdF ; mapCompose = mapComposeF } record { onObj = onObjG ; onArrow = onArrowG ; mapId = mapIdG ; mapCompose = mapComposeG } prObj prArr rewrite prObj = {!!}
VEC : {a : _} Functor (SET {a}) (SET {a})
VEC {a} n = record { onObj = λ X Vec X n ; onArrow = λ f λ v vecMap f v ; mapId = extensionality mapId' ; mapCompose = λ f g extensionality λ vec help f g vec }
where
vecMapLemma : {a : _} {T : Set a} {n : } (v : Vec T n) vecMap (Category.id SET T) v v
vecMapLemma {a} v with inspect (SET {a})
vecMapLemma {a} v | y with SetCopy = vecMapIdFact (λ i refl) v
mapId' : {a : _} {T : Set a} {n : } (v : Vec T n) vecMap (Category.id SET T) v Category.id SET (Vec T n) v
mapId' v rewrite vecMapLemma v = refl
help : {a n} {X Y Z : Category.objects (SET {a})} (f : X Y) (g : Y Z) (vec : Vec X n) vecMap (λ x g (f x)) vec vecMap g (vecMap f vec)
help f g vec = equalityCommutative (vecMapCompositionFact (λ x refl) vec)
CATEGORY : {a b : _} Category {lsuc b lsuc a} {b a}
CATEGORY {a} {b} = record { objects = Category {a} {b} ; arrows = λ C D Functor C D ; _∘_ = λ F G functorCompose F G ; id = λ C idFunctor C ; rightId = λ F {!!} ; leftId = λ F {!!} ; associative = {!!} }
where
rightIdFact : {a b c d : _} {C : Category {a} {b}} {D : Category {c} {d}} (F : Functor C D) functorCompose (idFunctor D) F F
rightIdFact {C = C} {D} F = {!!}

66
CyclicGroups.agda Normal file
View File

@@ -0,0 +1,66 @@
{-# OPTIONS --safe --warning=error #-}
open import LogicalFormulae
open import Setoids
open import Functions
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Naturals
open import Integers
open import FinSet
open import Groups
module CyclicGroups where
positiveEltPower : {m n : _} {A : Set m} {S : Setoid {m} {n} A} {_·_ : A A A} (G : Group S _·_) (x : A) (i : ) A
positiveEltPower G x 0 = Group.identity G
positiveEltPower {_·_ = _·_} G x (succ i) = x · (positiveEltPower G x i)
positiveEltPowerCollapse : {m n : _} {A : Set m} {S : Setoid {m} {n} A} {_+_ : A A A} (G : Group S _+_) (x : A) (i j : ) Setoid.__ S ((positiveEltPower G x i) + (positiveEltPower G x j)) (positiveEltPower G x (i +N j))
positiveEltPowerCollapse G x zero j = Group.multIdentLeft G
positiveEltPowerCollapse {S = S} G x (succ i) j = transitive (symmetric multAssoc) (wellDefined reflexive (positiveEltPowerCollapse G x i j))
where
open Setoid S
open Group G
open Transitive (Equivalence.transitiveEq eq)
open Symmetric (Equivalence.symmetricEq eq)
open Reflexive (Equivalence.reflexiveEq eq)
elementPower : {m n : _} {A : Set m} {S : Setoid {m} {n} A} {_·_ : A A A} (G : Group S _·_) (x : A) (i : ) A
elementPower G x (nonneg i) = positiveEltPower G x i
elementPower {_·_ = _·_} G x (negSucc i) = Group.inverse G (positiveEltPower G x (succ i))
record CyclicGroup {m n : _} {A : Set m} {S : Setoid {m} {n} A} {_·_ : A A A} (G : Group S _·_) : Set (m n) where
field
generator : A
cyclic : {a : A} (Sg (λ i Setoid.__ S (elementPower G generator i) a))
elementPowerCollapse : {m n : _} {A : Set m} {S : Setoid {m} {n} A} {_·_ : A A A} (G : Group S _·_) (x : A) (i j : ) Setoid.__ S ((elementPower G x i) · (elementPower G x j)) (elementPower G x (i +Z j))
elementPowerCollapse {S = S} {_+_} G x (nonneg a) (nonneg b) rewrite addingNonnegIsHom a b = positiveEltPowerCollapse G x a b
elementPowerCollapse G x (nonneg a) (negSucc b) = {!!}
elementPowerCollapse G x (negSucc a) (nonneg b) = {!!}
elementPowerCollapse G x (negSucc a) (negSucc b) = {!!}
elementPowerHom : {m n : _} {A : Set m} {S : Setoid {m} {n} A} {_·_ : A A A} (G : Group S _·_) (x : A) GroupHom Group G (λ i elementPower G x i)
GroupHom.groupHom (elementPowerHom {S = S} G x) {a} {b} = symmetric (elementPowerCollapse G x a b)
where
open Setoid S
open Group G
open Symmetric (Equivalence.symmetricEq eq)
GroupHom.wellDefined (elementPowerHom {S = S} G x) {.y} {y} refl = Reflexive.reflexive (Equivalence.reflexiveEq (Setoid.eq S))
subgroupOfCyclicIsCyclic : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+A_ : A A A} {_+B_ : B B B} {G : Group S _+A_} {H : Group T _+B_} {f : B A} {fHom : GroupHom H G f} Subgroup G H fHom CyclicGroup G CyclicGroup H
CyclicGroup.generator (subgroupOfCyclicIsCyclic {f = f} subgrp record { generator = generator ; cyclic = cyclic }) = {!f generator!}
CyclicGroup.cyclic (subgroupOfCyclicIsCyclic subgrp gCyclic) = {!!}
-- Prefer to prove that subgroup of cyclic is cyclic, then deduce this like with abelian groups
{-
cyclicIsGroupProperty : {m n o p : _} {A : Set m} {B : Set o} {S : Setoid {m} {n} A} {T : Setoid {o} {p} B} {_+_ : A → A → A} {_*_ : B → B → B} {G : Group S _+_} {H : Group T _*_} → GroupsIsomorphic G H → CyclicGroup G → CyclicGroup H
CyclicGroup.generator (cyclicIsGroupProperty {H = H} iso G) = GroupsIsomorphic.isomorphism iso (CyclicGroup.generator G)
CyclicGroup.cyclic (cyclicIsGroupProperty {H = H} iso G) {a} with GroupIso.surj (GroupsIsomorphic.proof iso) {a}
CyclicGroup.cyclic (cyclicIsGroupProperty {H = H} iso G) {a} | a' , b with CyclicGroup.cyclic G {a'}
... | pow , prPow = pow , {!!}
-}
-- Proof of abelianness of cyclic groups: a cyclic group is the image of elementPowerHom into Z, so is isomorphic to a subgroup of Z. All subgroups of an abelian group are abelian.
cyclicGroupIsAbelian : {m n : _} {A : Set m} {S : Setoid {m} {n} A} {_+_ : A A A} {G : Group S _+_} (cyclic : CyclicGroup G) AbelianGroup G
cyclicGroupIsAbelian cyclic = {!!}

33
Fields.agda Normal file
View File

@@ -0,0 +1,33 @@
{-# OPTIONS --safe --warning=error #-}
open import LogicalFormulae
open import Groups
open import Rings
open import Setoids
open import Orders
open import IntegralDomains
open import Functions
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
module Fields where
record Field {m n : _} {A : Set m} {S : Setoid {m} {n} A} {_+_ : A A A} {_*_ : A A A} (R : Ring S _+_ _*_) : Set (lsuc m n) where
open Ring R
open Group additiveGroup
open Setoid S
field
allInvertible : (a : A) ((a Group.identity (Ring.additiveGroup R)) False) Sg A (λ t t * a 1R)
nontrivial : (0R 1R) False
{-
record OrderedField {n} {A : Set n} {R : Ring A} (F : Field R) : Set (lsuc n) where
open Field F
field
ord : TotalOrder A
open TotalOrder ord
open Ring R
field
productPos : {a b : A} → (0R < a) → (0R < b) → (0R < (a * b))
orderRespectsAddition : {a b c : A} → (a < b) → (a + c) < (b + c)
-}

499
FinSet.agda Normal file
View File

@@ -0,0 +1,499 @@
{-# OPTIONS --safe --warning=error #-}
open import Functions
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import LogicalFormulae
open import Naturals
open import Orders
module FinSet where
data FinSet : (n : ) Set where
fzero : {n : } FinSet (succ n)
fsucc : {n : } FinSet n FinSet (succ n)
data FinNotEquals : {n : } (a b : FinSet (succ n)) Set where
fne2 : (a b : FinSet 2) ((a fzero) && (b fsucc (fzero))) || ((b fzero) && (a fsucc (fzero))) FinNotEquals {1} a b
fneN : {n : } (a b : FinSet (succ (succ (succ n)))) (((a fzero) && (Sg (FinSet (succ (succ n))) (λ c b fsucc c))) || ((Sg (FinSet (succ (succ n))) (λ c a fsucc c)) && (b fzero))) || (Sg (FinSet (succ (succ n)) && FinSet (succ (succ n))) (λ t (a fsucc (_&&_.fst t)) & (b fsucc (_&&_.snd t)) & FinNotEquals (_&&_.fst t) (_&&_.snd t))) FinNotEquals {succ (succ n)} a b
fsuccInjective : {n : } {a b : FinSet n} fsucc a fsucc b a b
fsuccInjective refl = refl
fneSymmetric : {n : } {a b : FinSet (succ n)} FinNotEquals a b FinNotEquals b a
fneSymmetric {.1} {.a1} {.b1} (fne2 a1 b1 (inl x)) = fne2 b1 a1 (inr x)
fneSymmetric {.1} {.a1} {.b1} (fne2 a1 b1 (inr x)) = fne2 b1 a1 (inl x)
fneSymmetric {.(succ (succ _))} {.fzero} {b} (fneN .fzero b (inl (inl (refl ,, snd)))) = fneN b fzero (inl (inr (snd ,, refl)))
fneSymmetric {.(succ (succ _))} {a} {.fzero} (fneN a .fzero (inl (inr (fst ,, refl)))) = fneN fzero a (inl (inl (refl ,, fst)))
fneSymmetric {.(succ (succ _))} {a} {b} (fneN a b (inr ((fst ,, snd) , record { one = o ; two = p ; three = q }))) = fneN b a (inr ((snd ,, fst) , record { one = p ; two = o ; three = fneSymmetric q }))
underlyingProof : {l m : _} {L : Set l} {pr : L Set m} (a : Sg L pr) pr (underlying a)
underlyingProof (a , b) = b
sgEq : {l m : _} {L : Set l} {pr : L Set m} {a b : Sg L pr} (underlying a underlying b) ({c : L} (r s : pr c) r s) (a b)
sgEq {l} {m} {L} {prop} {(a , b1)} {(.a , b)} refl pr2 with pr2 {a} b b1
sgEq {l} {m} {L} {prop} {(a , b1)} {(.a , .b1)} refl pr2 | refl = refl
finNotEqualsRefl : {n : } {a b : FinSet (succ n)} (p1 p2 : FinNotEquals a b) p1 p2
finNotEqualsRefl {.1} {.fzero} {.(fsucc fzero)} (fne2 .fzero .(fsucc fzero) (inl (refl ,, refl))) (fne2 .fzero .(fsucc fzero) (inl (refl ,, refl))) = refl
finNotEqualsRefl {.1} {.fzero} {.(fsucc fzero)} (fne2 .fzero .(fsucc fzero) (inl (refl ,, refl))) (fne2 .fzero .(fsucc fzero) (inr (() ,, snd)))
finNotEqualsRefl {.1} {.(fsucc fzero)} {.fzero} (fne2 .(fsucc fzero) .fzero (inr (refl ,, refl))) (fne2 .(fsucc fzero) .fzero (inl (() ,, snd)))
finNotEqualsRefl {.1} {.(fsucc fzero)} {.fzero} (fne2 .(fsucc fzero) .fzero (inr (refl ,, refl))) (fne2 .(fsucc fzero) .fzero (inr (refl ,, refl))) = refl
finNotEqualsRefl {.(succ (succ _))} {.fzero} {.(fsucc a)} (fneN .fzero .(fsucc a) (inl (inl (refl ,, (a , refl))))) (fneN .fzero .(fsucc a) (inl (inl (refl ,, (.a , refl))))) = refl
finNotEqualsRefl {.(succ (succ _))} {.fzero} {.(fsucc a)} (fneN .fzero .(fsucc a) (inl (inl (refl ,, (a , refl))))) (fneN .fzero .(fsucc a) (inl (inr ((a₁ , ()) ,, snd))))
finNotEqualsRefl {.(succ (succ _))} {.fzero} {.(fsucc a)} (fneN .fzero .(fsucc a) (inl (inl (refl ,, (a , refl))))) (fneN .fzero .(fsucc a) (inr ((fst ,, snd) , b))) = exFalso q
where
p : fzero fsucc fst
p = _&_&_.one b
q : False
q with p
... | ()
finNotEqualsRefl {.(succ (succ _))} {.(fsucc a)} {.fzero} (fneN .(fsucc a) .fzero (inl (inr ((a , refl) ,, refl)))) (fneN .(fsucc a) .fzero (inl (inl (() ,, snd))))
finNotEqualsRefl {.(succ (succ _))} {.(fsucc a)} {.fzero} (fneN .(fsucc a) .fzero (inl (inr ((a , refl) ,, refl)))) (fneN .(fsucc a) .fzero (inl (inr ((.a , refl) ,, refl)))) = refl
finNotEqualsRefl {.(succ (succ _))} {.(fsucc a)} {.fzero} (fneN .(fsucc a) .fzero (inl (inr ((a , refl) ,, refl)))) (fneN .(fsucc a) .fzero (inr ((fst ,, snd) , b))) = exFalso q
where
p : fzero fsucc snd
p = _&_&_.two b
q : False
q with p
... | ()
finNotEqualsRefl {.(succ (succ _))} {a} {b} (fneN a b (inr (record { fst = fst ; snd = snd₁ } , snd))) (fneN .a .b (inl (inl (fst1 ,, snd₂)))) = exFalso q
where
p : fzero fsucc fst
p = transitivity (equalityCommutative fst1) (_&_&_.one snd)
q : False
q with p
... | ()
finNotEqualsRefl {.(succ (succ _))} {a} {b} (fneN a b (inr (record { fst = fst ; snd = snd1 } , snd))) (fneN .a .b (inl (inr (fst₁ ,, snd2)))) = exFalso q
where
p : fzero fsucc snd1
p = transitivity (equalityCommutative snd2) (_&_&_.two snd)
q : False
q with p
... | ()
finNotEqualsRefl {.(succ (succ _))} {.fzero} {b} (fneN fzero b (inr (record { fst = fst ; snd = snd₁ } , snd))) (fneN .fzero .b (inr ((fst1 ,, snd₂) , b1))) = exFalso q
where
p : fzero fsucc fst1
p = _&_&_.one b1
q : False
q with p
... | ()
finNotEqualsRefl {.(succ (succ _))} {.(fsucc a)} {.fzero} (fneN (fsucc a) fzero (inr (record { fst = fst ; snd = snd₁ } , snd))) (fneN .(fsucc a) .fzero (inr ((fst₁ ,, snd2) , b1))) = exFalso q
where
p : fzero fsucc snd2
p = _&_&_.two b1
q : False
q with p
... | ()
finNotEqualsRefl {.(succ (succ _))} {.(fsucc a)} {.(fsucc b)} (fneN (fsucc a) (fsucc b) (inr (record { fst = fst ; snd = snd1 } , snd))) (fneN .(fsucc a) .(fsucc b) (inr ((fst1 ,, snd2) , b1))) = ans
where
t : a fst1
t = fsuccInjective (_&_&_.one b1)
t' : a fst
t' = fsuccInjective (_&_&_.one snd)
u : b snd1
u = fsuccInjective (_&_&_.two snd)
u' : b snd2
u' = fsuccInjective (_&_&_.two b1)
equality : {c : FinSet (succ (succ _)) && FinSet (succ (succ _))} (r1 s : (fsucc a fsucc (_&&_.fst c)) & (fsucc b fsucc (_&&_.snd c)) & FinNotEquals (_&&_.fst c) (_&&_.snd c)) r1 s
equality record { one = refl ; two = refl ; three = q } record { one = refl ; two = refl ; three = q' } = applyEquality (λ t record { one = refl ; two = refl ; three = t }) (finNotEqualsRefl q q')
r : (fst ,, snd1) (fst1 ,, snd2)
r rewrite equalityCommutative t | equalityCommutative t' | equalityCommutative u | equalityCommutative u' = refl
ans : fneN (fsucc a) (fsucc b) (inr ((fst ,, snd1) , snd)) fneN (fsucc a) (fsucc b) (inr ((fst1 ,, snd2) , b1))
ans = applyEquality (λ t fneN (fsucc a) (fsucc b) (inr t)) (sgEq r equality)
finSetNotEquals : {n : } {a b : FinSet (succ n)} (a b False) FinNotEquals {n} a b
finSetNotEquals {zero} {fzero} {fzero} pr = exFalso (pr refl)
finSetNotEquals {zero} {fzero} {fsucc ()} pr
finSetNotEquals {zero} {fsucc ()} {b} pr
finSetNotEquals {succ zero} {fzero} {fzero} pr = exFalso (pr refl)
finSetNotEquals {succ zero} {fzero} {fsucc fzero} pr = fne2 fzero (fsucc fzero) (inl (refl ,, refl))
finSetNotEquals {succ zero} {fzero} {fsucc (fsucc ())} pr
finSetNotEquals {succ zero} {fsucc fzero} {fzero} pr = fne2 (fsucc fzero) fzero (inr (refl ,, refl))
finSetNotEquals {succ zero} {fsucc fzero} {fsucc fzero} pr = exFalso (pr refl)
finSetNotEquals {succ zero} {fsucc fzero} {fsucc (fsucc ())} pr
finSetNotEquals {succ zero} {fsucc (fsucc ())}
finSetNotEquals {succ (succ n)} {fzero} {fzero} pr = exFalso (pr refl)
finSetNotEquals {succ (succ n)} {fzero} {fsucc b} pr = fneN fzero (fsucc b) (inl (inl (refl ,, (b , refl))))
finSetNotEquals {succ (succ n)} {fsucc a} {fzero} pr = fneN (fsucc a) fzero (inl (inr ((a , refl) ,, refl)))
finSetNotEquals {succ (succ n)} {fsucc a} {fsucc b} pr = fneN (fsucc a) (fsucc b) (inr ans)
where
q : a b False
q refl = pr refl
t : FinNotEquals {succ n} a b
t = finSetNotEquals {succ n} {a} {b} q
ans : Sg (FinSet (succ (succ n)) && FinSet (succ (succ n))) (λ x (fsucc a fsucc (_&&_.fst x)) & fsucc b fsucc (_&&_.snd x) & FinNotEquals (_&&_.fst x) (_&&_.snd x))
ans with t
ans | fne2 .fzero .(fsucc fzero) (inl (refl ,, refl)) = (a ,, b) , record { one = refl ; two = refl ; three = fne2 fzero (fsucc fzero) (inl (refl ,, refl)) }
ans | fne2 .(fsucc fzero) .fzero (inr (refl ,, refl)) = (a ,, b) , record { one = refl ; two = refl ; three = fne2 (fsucc fzero) fzero (inr (refl ,, refl)) }
ans | fneN a b _ = (a ,, b) , record { one = refl ; two = refl ; three = t }
fzeroNotFsucc : {n : } {a : _} fzero fsucc {succ n} a False
fzeroNotFsucc ()
finSetNotEqualsSame : {n : } {a : FinSet (succ n)} FinNotEquals a a False
finSetNotEqualsSame {.1} {fzero} (fne2 .fzero .fzero (inl (fst ,, ())))
finSetNotEqualsSame {.1} {fzero} (fne2 .fzero .fzero (inr (fst ,, ())))
finSetNotEqualsSame {.(succ (succ _))} {fzero} (fneN .fzero .fzero (inl (inl (refl ,, (a , ())))))
finSetNotEqualsSame {.(succ (succ _))} {fzero} (fneN .fzero .fzero (inl (inr ((a , ()) ,, refl))))
finSetNotEqualsSame {.(succ (succ _))} {fzero} (fneN .fzero .fzero (inr ((fst ,, snd) , record { one = () ; two = p ; three = q })))
finSetNotEqualsSame {.1} {fsucc a} (fne2 .(fsucc a) .(fsucc a) (inl (() ,, snd)))
finSetNotEqualsSame {.1} {fsucc a} (fne2 .(fsucc a) .(fsucc a) (inr (() ,, snd)))
finSetNotEqualsSame {.(succ (succ _))} {fsucc a} (fneN .(fsucc a) .(fsucc a) (inl (inl (() ,, snd))))
finSetNotEqualsSame {.(succ (succ _))} {fsucc a} (fneN .(fsucc a) .(fsucc a) (inl (inr (fst ,, ()))))
finSetNotEqualsSame {.(succ (succ _))} {fsucc a} (fneN .(fsucc a) .(fsucc a) (inr ((.a ,, .a) , record { one = refl ; two = refl ; three = q }))) = finSetNotEqualsSame q
finNotEqualsFsucc : {n : } {a b : FinSet (succ n)} FinNotEquals (fsucc a) (fsucc b) FinNotEquals a b
finNotEqualsFsucc {.0} {a} {b} (fne2 .(fsucc a) .(fsucc b) (inl (() ,, snd)))
finNotEqualsFsucc {.0} {a} {b} (fne2 .(fsucc a) .(fsucc b) (inr (() ,, snd)))
finNotEqualsFsucc {n} {a} {b} (fneN .(fsucc a) .(fsucc b) (inl (inl (() ,, snd))))
finNotEqualsFsucc {n} {a} {b} (fneN .(fsucc a) .(fsucc b) (inl (inr (fst ,, ()))))
finNotEqualsFsucc {n} {a} {b} (fneN .(fsucc a) .(fsucc b) (inr ((fst ,, snd) , prf))) = identityOfIndiscernablesRight _ _ _ FinNotEquals ans (equalityCommutative b=snd)
where
a=fst : a fst
a=fst = fsuccInjective (_&_&_.one prf)
b=snd : b snd
b=snd = fsuccInjective (_&_&_.two prf)
ans : FinNotEquals a snd
ans = identityOfIndiscernablesLeft _ _ _ FinNotEquals (_&_&_.three prf) (equalityCommutative a=fst)
finSetNotEquals' : {n : } {a b : FinSet (succ n)} FinNotEquals a b (a b False)
finSetNotEquals' {n} {a} {.a} fne refl = finSetNotEqualsSame fne
finset0Empty : FinSet zero False
finset0Empty ()
finset1OnlyOne : (a b : FinSet 1) a b
finset1OnlyOne fzero fzero = refl
finset1OnlyOne fzero (fsucc ())
finset1OnlyOne (fsucc ()) b
intoSmaller : {n m : } (n <N m) (FinSet n FinSet m)
intoSmaller {zero} {m} n<m = t
where
t : FinSet 0 FinSet m
t ()
intoSmaller {succ n} {zero} ()
intoSmaller {succ n} {succ m} n<m with intoSmaller (canRemoveSuccFrom<N n<m)
intoSmaller {succ n} {succ m} n<m | prev = t
where
t : FinSet (succ n) FinSet (succ m)
t fzero = fzero
t (fsucc arg) = fsucc (prev arg)
intoSmallerInj : {n m : } (n<m : n <N m) Injection (intoSmaller n<m)
intoSmallerInj {zero} {zero} (le x ())
intoSmallerInj {zero} {succ m} n<m = record { property = inj }
where
t : FinSet 0 FinSet (succ m)
t ()
inj : {x y : FinSet zero} intoSmaller (succIsPositive m) x intoSmaller (succIsPositive m) y x y
inj {()} {y}
intoSmallerInj {succ n} {zero} ()
intoSmallerInj {succ n} {succ m} n<m with intoSmallerInj (canRemoveSuccFrom<N n<m)
intoSmallerInj {succ n} {succ m} n<m | prevInj = inj
where
inj : Injection (intoSmaller n<m)
Injection.property inj {fzero} {fzero} pr = refl
Injection.property inj {fzero} {fsucc y} ()
Injection.property inj {fsucc x} {fzero} ()
Injection.property inj {fsucc x} {fsucc y} pr = applyEquality fsucc (Injection.property prevInj (fsuccInjective pr))
toNat : {n : } (a : FinSet n)
toNat {.(succ _)} fzero = 0
toNat {.(succ _)} (fsucc a) = succ (toNat a)
toNatSmaller : {n : } (a : FinSet n) toNat a <N n
toNatSmaller {zero} ()
toNatSmaller {succ n} fzero = succIsPositive n
toNatSmaller {succ n} (fsucc a) = succPreservesInequality (toNatSmaller a)
ofNat : {n : } (m : ) (m <N n) FinSet n
ofNat {zero} zero (le x ())
ofNat {succ n} zero m<n = fzero
ofNat {zero} (succ m) (le x ())
ofNat {succ n} (succ m) m<n = fsucc (ofNat {n} m (canRemoveSuccFrom<N m<n))
ofNatInjective : {n : } (x y : ) (pr : x <N n) (pr2 : y <N n) ofNat x pr ofNat y pr2 x y
ofNatInjective {zero} zero zero () y<n pr
ofNatInjective {zero} zero (succ y) () y<n pr
ofNatInjective {zero} (succ x) zero x<n () pr
ofNatInjective {zero} (succ x) (succ y) () y<n pr
ofNatInjective {succ n} zero zero x<n y<n pr = refl
ofNatInjective {succ n} zero (succ y) x<n y<n ()
ofNatInjective {succ n} (succ x) zero x<n y<n ()
ofNatInjective {succ n} (succ x) (succ y) x<n y<n pr = applyEquality succ (ofNatInjective x y (canRemoveSuccFrom<N x<n) (canRemoveSuccFrom<N y<n) (fsuccInjective pr))
toNatOfNat : {n : } (a : ) (a<n : a <N n) toNat (ofNat a a<n) a
toNatOfNat {zero} zero (le x ())
toNatOfNat {zero} (succ a) (le x ())
toNatOfNat {succ n} zero a<n = refl
toNatOfNat {succ n} (succ a) a<n = applyEquality succ (toNatOfNat a (canRemoveSuccFrom<N a<n))
ofNatToNat : {n : } (a : FinSet n) ofNat (toNat a) (toNatSmaller a) a
ofNatToNat {zero} ()
ofNatToNat {succ n} fzero = refl
ofNatToNat {succ n} (fsucc a) = applyEquality fsucc indHyp
where
indHyp : ofNat (toNat a) (canRemoveSuccFrom<N (succPreservesInequality (toNatSmaller a))) a
indHyp rewrite <NRefl (canRemoveSuccFrom<N (succPreservesInequality (toNatSmaller a))) (toNatSmaller a) = ofNatToNat a
intoSmallerLemm : {n m : } {n<m : n <N (succ m)} {m<sm : m <N succ m} (b : FinSet n) intoSmaller n<m b ofNat m m<sm False
intoSmallerLemm {.(succ _)} {m} {n<m} fzero pr with intoSmaller (canRemoveSuccFrom<N n<m)
intoSmallerLemm {.(succ _)} {zero} {n<m} fzero refl | bl = zeroNeverGreater (canRemoveSuccFrom<N n<m)
intoSmallerLemm {.(succ _)} {succ m} {n<m} fzero () | bl
intoSmallerLemm {.(succ _)} {m} {n<m} (fsucc b) pr with inspect (intoSmaller (canRemoveSuccFrom<N n<m))
intoSmallerLemm {.(succ _)} {zero} {n<m} (fsucc b) pr | bl with _ = zeroNeverGreater (canRemoveSuccFrom<N n<m)
intoSmallerLemm {.(succ _)} {succ m} {n<m} {m<sm} (fsucc b) pr | bl with p = intoSmallerLemm {m<sm = canRemoveSuccFrom<N m<sm} b (fsuccInjective pr)
intoSmallerNotSurj : {n m : } {n<m : n <N m} Surjection (intoSmaller n<m) False
intoSmallerNotSurj {n} {zero} {le x ()} record { property = property }
intoSmallerNotSurj {zero} {succ zero} {n<m} record { property = property } with property fzero
... | () , _
intoSmallerNotSurj {succ n} {succ zero} {n<m} record { property = property } = zeroNeverGreater (canRemoveSuccFrom<N n<m)
intoSmallerNotSurj {0} {succ (succ m)} {n<m} record { property = property } with property fzero
... | () , _
intoSmallerNotSurj {succ n} {succ (succ m)} {n<m} record { property = property } = problem
where
notHit : FinSet (succ (succ m))
notHit = ofNat (succ m) (le zero refl)
hitting : Sg (FinSet (succ n)) (λ i intoSmaller n<m i notHit)
hitting = property notHit
problem : False
problem with hitting
... | a , pr = intoSmallerLemm {succ n} {succ m} {n<m} {le 0 refl} a pr
finsetDecidableEquality : {n : } (x y : FinSet n) (x y) || ((x y) False)
finsetDecidableEquality fzero fzero = inl refl
finsetDecidableEquality fzero (fsucc y) = inr λ ()
finsetDecidableEquality (fsucc x) fzero = inr λ ()
finsetDecidableEquality (fsucc x) (fsucc y) with finsetDecidableEquality x y
finsetDecidableEquality (fsucc x) (fsucc y) | inl pr rewrite pr = inl refl
finsetDecidableEquality (fsucc x) (fsucc y) | inr pr = inr (λ f pr (fsuccInjective f))
subInjection : {n : } {f : FinSet (succ (succ n)) FinSet (succ (succ n))} Injection f (FinSet (succ n) FinSet (succ n))
subInjection {n} {f} inj x with inspect (f (fsucc x))
subInjection {f = f} inj x | fzero with _ with inspect (f fzero)
subInjection {f = f} inj x | fzero with pr | fzero with pr2 = exFalso (fzeroNotFsucc (Injection.property inj (transitivity pr2 (equalityCommutative pr))))
subInjection {f = f} inj x | fzero with _ | fsucc bl with _ = bl
subInjection {f = f} inj x | fsucc bl with _ = bl
subInjIsInjective : {n : } {f : FinSet (succ (succ n)) FinSet (succ (succ n))} (inj : Injection f) Injection (subInjection inj)
Injection.property (subInjIsInjective {f = f} inj) {x} {y} pr with inspect (f (fsucc x))
Injection.property (subInjIsInjective {f = f} inj) {x} {y} pr | fzero with _ with inspect (f (fzero))
Injection.property (subInjIsInjective {f = f} inj) {x} {y} pr | fzero with pr1 | fzero with pr2 = exFalso (fzeroNotFsucc (Injection.property inj (transitivity pr2 (equalityCommutative pr1))))
Injection.property (subInjIsInjective {f = f} inj) {x} {y} pr | fzero with pr1 | fsucc bl with _ with inspect (f (fsucc y))
Injection.property (subInjIsInjective {f = f} inj) {x} {y} pr | fzero with pr1 | fsucc bl with _ | fzero with x₁ with inspect (f (fzero))
Injection.property (subInjIsInjective {f = f} inj) {x} {y} pr | fzero with pr1 | fsucc bl with _ | fzero with x1 | fzero with x2 = exFalso (fzeroNotFsucc (Injection.property inj (transitivity x2 (equalityCommutative x1))))
Injection.property (subInjIsInjective {f = f} inj) {x} {y} refl | fzero with pr1 | fsucc .bl2 with _ | fzero with x1 | fsucc bl2 with _ = fsuccInjective (Injection.property inj (transitivity pr1 (equalityCommutative x1)))
Injection.property (subInjIsInjective {f = f} inj) {x} {y} refl | fzero with pr1 | fsucc .bl2 with pr2 | fsucc bl2 with pr3 = exFalso (fzeroNotFsucc (Injection.property inj (transitivity pr2 (equalityCommutative pr3))))
Injection.property (subInjIsInjective {f = f} inj) {x} {y} pr | fsucc bl with _ with inspect (f (fsucc y))
Injection.property (subInjIsInjective {f = f} inj) {x} {y} pr | fsucc bl with _ | fzero with x₁ with inspect (f fzero)
Injection.property (subInjIsInjective {f = f} inj) {x} {y} pr | fsucc bl with _ | fzero with x1 | fzero with x2 = exFalso (fzeroNotFsucc (Injection.property inj (transitivity x2 (equalityCommutative x1))))
Injection.property (subInjIsInjective {f = f} inj) {x} {y} refl | fsucc .bl2 with x1 | fzero with x₁ | fsucc bl2 with x2 = exFalso (fzeroNotFsucc (Injection.property inj (transitivity x2 (equalityCommutative x1))))
Injection.property (subInjIsInjective {f = f} inj) {x} {y} refl | fsucc .y1 with pr1 | fsucc y1 with pr2 = fsuccInjective (Injection.property inj (transitivity pr1 (equalityCommutative pr2)))
onepointBij : (f : FinSet 1 FinSet 1) Bijection f
Injection.property (Bijection.inj (onepointBij f)) {x} {y} _ = finset1OnlyOne x y
Surjection.property (Bijection.surj (onepointBij f)) fzero with inspect (f fzero)
... | fzero with pr = fzero , pr
... | fsucc () with _
Surjection.property (Bijection.surj (onepointBij f)) (fsucc ())
nopointBij : (f : FinSet 0 FinSet 0) Bijection f
Injection.property (Bijection.inj (nopointBij f)) {()}
Surjection.property (Bijection.surj (nopointBij f)) ()
flip : {n : } (f : FinSet (succ n) FinSet (succ n)) FinSet (succ n) FinSet (succ n)
flip {n} f fzero = f fzero
flip {n} f (fsucc a) with inspect (f fzero)
flip {n} f (fsucc a) | fzero with x = fsucc a
flip {n} f (fsucc a) | fsucc y with x with finsetDecidableEquality a y
flip {n} f (fsucc a) | fsucc y with x | inl a=y = fzero
flip {n} f (fsucc a) | fsucc y with x | inr a!=y = fsucc a
flipSwapsF0 : {n : } {f : FinSet (succ n) FinSet (succ n)} (flip f (f fzero)) fzero
flipSwapsF0 {zero} {f} with inspect (f fzero)
flipSwapsF0 {zero} {f} | fzero with x rewrite x = x
flipSwapsF0 {zero} {f} | fsucc () with x
flipSwapsF0 {succ n} {f = f} with inspect (f fzero)
flipSwapsF0 {succ n} {f = f} | fzero with pr rewrite pr = pr
flipSwapsF0 {succ n} {f = f} | fsucc b with pr rewrite pr = ans
where
ans : flip f (fsucc b) fzero
ans with inspect (f fzero)
ans | fzero with x = exFalso (fzeroNotFsucc (transitivity (equalityCommutative x) pr))
ans | fsucc y with x with finsetDecidableEquality b y
ans | fsucc y with x | inl b=y = refl
ans | fsucc y with x | inr b!=y = exFalso (b!=y (fsuccInjective (transitivity (equalityCommutative pr) x)))
flipSwapsZero : {n : } {f : FinSet (succ n) FinSet (succ n)} (flip f fzero) f fzero
flipSwapsZero {n} {f} = refl
flipMaintainsEverythingElse : {n : } {f : FinSet (succ n) FinSet (succ n)} (x : FinSet n) ((fsucc x f fzero) False) flip f (fsucc x) fsucc x
flipMaintainsEverythingElse {n} {f} x x!=f0 with inspect (f fzero)
flipMaintainsEverythingElse {n} {f} x x!=f0 | fzero with pr = refl
flipMaintainsEverythingElse {n} {f} x x!=f0 | fsucc bl with pr with finsetDecidableEquality x bl
flipMaintainsEverythingElse {n} {f} x x!=f0 | fsucc bl with pr | inl x=bl = exFalso (x!=f0 (transitivity (applyEquality fsucc x=bl) (equalityCommutative pr)))
flipMaintainsEverythingElse {n} {f} x x!=f0 | fsucc bl with pr | inr x!=bl = refl
bijFlip : {n : } {f : FinSet (succ n) FinSet (succ n)} Bijection (flip f)
bijFlip {zero} {f} = onepointBij (flip f)
Injection.property (Bijection.inj (bijFlip {succ n} {f})) {fzero} {fzero} flx=fly = refl
Injection.property (Bijection.inj (bijFlip {succ n} {f})) {fzero} {fsucc y} flx=fly with inspect (f fzero)
Injection.property (Bijection.inj (bijFlip {succ n} {f})) {fzero} {fsucc y} flx=fly | fzero with x rewrite x = exFalso (fzeroNotFsucc flx=fly)
Injection.property (Bijection.inj (bijFlip {succ n} {f})) {fzero} {fsucc y} flx=fly | fsucc f0 with x with finsetDecidableEquality y f0
Injection.property (Bijection.inj (bijFlip {succ n} {f})) {fzero} {fsucc y} flx=fly | fsucc f0 with x | inl x₁ = exFalso (fzeroNotFsucc (transitivity (equalityCommutative flx=fly) x))
Injection.property (Bijection.inj (bijFlip {succ n} {f})) {fzero} {fsucc y} flx=fly | fsucc f0 with x | inr pr = exFalso (pr (fsuccInjective (transitivity (equalityCommutative flx=fly) x)))
Injection.property (Bijection.inj (bijFlip {succ n} {f})) {fsucc x} {fzero} flx=fly with inspect (f fzero)
Injection.property (Bijection.inj (bijFlip {succ n} {f})) {fsucc x} {fzero} flx=fly | fzero with pr = transitivity flx=fly pr
Injection.property (Bijection.inj (bijFlip {succ n} {f})) {fsucc x} {fzero} flx=fly | fsucc f0 with x₁ with finsetDecidableEquality x f0
Injection.property (Bijection.inj (bijFlip {succ n} {f})) {fsucc x} {fzero} flx=fly | fsucc f0 with pr | inl x₂ = exFalso (fzeroNotFsucc (transitivity flx=fly pr))
Injection.property (Bijection.inj (bijFlip {succ n} {f})) {fsucc x} {fzero} flx=fly | fsucc f0 with x1 | inr x!=f0 = exFalso (x!=f0 (fsuccInjective (transitivity flx=fly x1)))
Injection.property (Bijection.inj (bijFlip {succ n} {f})) {fsucc x} {fsucc y} flx=fly with inspect (f fzero)
Injection.property (Bijection.inj (bijFlip {succ n} {f})) {fsucc x} {fsucc y} flx=fly | fzero with x₁ = flx=fly
Injection.property (Bijection.inj (bijFlip {succ n} {f})) {fsucc x} {fsucc y} flx=fly | fsucc f0 with x₁ with finsetDecidableEquality y f0
Injection.property (Bijection.inj (bijFlip {succ n} {f})) {fsucc x} {fsucc y} flx=fly | fsucc f0 with x₁ | inl y=f0 with finsetDecidableEquality x f0
Injection.property (Bijection.inj (bijFlip {succ n} {f})) {fsucc x} {fsucc y} flx=fly | fsucc f0 with x₁ | inl y=f0 | inl x=f0 = applyEquality fsucc (transitivity x=f0 (equalityCommutative y=f0))
Injection.property (Bijection.inj (bijFlip {succ n} {f})) {fsucc x} {fsucc y} () | fsucc f0 with x₁ | inl y=f0 | inr x!=f0
Injection.property (Bijection.inj (bijFlip {succ n} {f})) {fsucc x} {fsucc y} flx=fly | fsucc f0 with x₁ | inr y!=f0 with finsetDecidableEquality x f0
Injection.property (Bijection.inj (bijFlip {succ n} {f})) {fsucc x} {fsucc y} () | fsucc f0 with x₁ | inr y!=f0 | inl x=f0
Injection.property (Bijection.inj (bijFlip {succ n} {f})) {fsucc x} {fsucc y} flx=fly | fsucc f0 with x₁ | inr y!=f0 | inr x!=f0 = flx=fly
Surjection.property (Bijection.surj (bijFlip {succ n} {f})) fzero = f fzero , flipSwapsF0 {f = f}
Surjection.property (Bijection.surj (bijFlip {succ n} {f})) (fsucc b) with inspect (f fzero)
Surjection.property (Bijection.surj (bijFlip {succ n} {f})) (fsucc b) | fzero with x = fsucc b , flipMaintainsEverythingElse {f = f} b λ pr fzeroNotFsucc (transitivity (equalityCommutative x) (equalityCommutative pr))
Surjection.property (Bijection.surj (bijFlip {succ n} {f})) (fsucc b) | fsucc y with x with finsetDecidableEquality y b
Surjection.property (Bijection.surj (bijFlip {succ n} {f})) (fsucc b) | fsucc y with x | inl y=b = fzero , transitivity x (applyEquality fsucc y=b)
Surjection.property (Bijection.surj (bijFlip {succ n} {f})) (fsucc b) | fsucc y with x | inr y!=b = fsucc b , flipMaintainsEverythingElse {f = f} b λ pr y!=b (fsuccInjective (transitivity (equalityCommutative x) (equalityCommutative pr)))
injectionIsBijectionFinset : {n : } {f : FinSet n FinSet n} Injection f Bijection f
injectionIsBijectionFinset {zero} {f} inj = record { inj = inj ; surj = record { property = λ () } }
injectionIsBijectionFinset {succ zero} {f} inj = record { inj = inj ; surj = record { property = ans } }
where
ans : (b : FinSet (succ zero)) Sg (FinSet (succ zero)) (λ a f a b)
ans fzero with inspect (f fzero)
ans fzero | fzero with x = fzero , x
ans fzero | fsucc () with x
ans (fsucc ())
injectionIsBijectionFinset {succ (succ n)} {f} inj with injectionIsBijectionFinset {succ n} (subInjIsInjective inj)
... | sb = ans
where
subSurj : Surjection (subInjection inj)
subSurj = Bijection.surj sb
tweakedF : FinSet (succ (succ n)) FinSet (succ (succ n))
tweakedF fzero = fzero
tweakedF (fsucc x) = fsucc (subInjection inj x)
tweakedBij : Bijection tweakedF
Injection.property (Bijection.inj tweakedBij) {fzero} {fzero} f'x=f'y = refl
Injection.property (Bijection.inj tweakedBij) {fzero} {fsucc y} ()
Injection.property (Bijection.inj tweakedBij) {fsucc x} {fzero} ()
Injection.property (Bijection.inj tweakedBij) {fsucc x} {fsucc y} f'x=f'y = applyEquality fsucc (Injection.property (subInjIsInjective inj) (fsuccInjective f'x=f'y))
Surjection.property (Bijection.surj tweakedBij) fzero = fzero , refl
Surjection.property (Bijection.surj tweakedBij) (fsucc b) with Surjection.property subSurj b
... | ans , prop = fsucc ans , applyEquality fsucc prop
compBij : Bijection (λ i flip f (tweakedF i))
compBij = bijectionComp tweakedBij bijFlip
undoTweakMakesF : {x : FinSet (succ (succ n))} flip f (tweakedF x) f x
undoTweakMakesF {fzero} = refl
undoTweakMakesF {fsucc x} with inspect (f fzero)
undoTweakMakesF {fsucc x} | fzero with x₁ with inspect (f (fsucc x))
undoTweakMakesF {fsucc x} | fzero with x₁ | fzero with x₂ with inspect (f fzero)
undoTweakMakesF {fsucc x} | fzero with x₁ | fzero with x2 | fzero with x3 = exFalso (fzeroNotFsucc (Injection.property inj (transitivity x3 (equalityCommutative x2))))
undoTweakMakesF {fsucc x} | fzero with x1 | fzero with x2 | fsucc y with x3 = exFalso (fzeroNotFsucc (Injection.property inj (transitivity x1 (equalityCommutative x2))))
undoTweakMakesF {fsucc x} | fzero with x1 | fsucc y with x2 = equalityCommutative x2
undoTweakMakesF {fsucc x} | fsucc y with x₁ with inspect (f (fsucc x))
undoTweakMakesF {fsucc x} | fsucc y with x₁ | fzero with x₂ with inspect (f fzero)
undoTweakMakesF {fsucc x} | fsucc y with x₁ | fzero with x2 | fzero with x3 = exFalso (fzeroNotFsucc (Injection.property inj (transitivity x3 (equalityCommutative x2))))
undoTweakMakesF {fsucc x} | fsucc y with x₁ | fzero with x₂ | fsucc pr with x₃ with finsetDecidableEquality pr y
undoTweakMakesF {fsucc x} | fsucc y with x₁ | fzero with x2 | fsucc pr with x₃ | inl x₄ = equalityCommutative x2
undoTweakMakesF {fsucc x} | fsucc y with x1 | fzero with x₂ | fsucc pr with x3 | inr pr!=y = exFalso (pr!=y (fsuccInjective (transitivity (equalityCommutative x3) x1)))
undoTweakMakesF {fsucc x} | fsucc y with x₁ | fsucc thi with x₂ with finsetDecidableEquality thi y
undoTweakMakesF {fsucc x} | fsucc .thi with x1 | fsucc thi with x2 | inl refl = exFalso false
where
p : f fzero f (fsucc x)
p = transitivity x1 (equalityCommutative x2)
q : fzero fsucc x
q = Injection.property inj p
false : False
false = fzeroNotFsucc q
undoTweakMakesF {fsucc x} | fsucc y with x₁ | fsucc thi with x2 | inr thi!=y = equalityCommutative x2
ans : Bijection f
ans = bijectionPreservedUnderExtensionalEq compBij undoTweakMakesF
pigeonhole : {n m : } (m <N n) {f : FinSet n FinSet m} Injection f False
pigeonhole {zero} {zero} () {f} record { property = property }
pigeonhole {zero} {succ m} () {f} record { property = property }
pigeonhole {succ n} {zero} m<n {f} record { property = property } = finset0Empty (f fzero)
pigeonhole {succ zero} {succ m} m<n {f} record { property = property } with canRemoveSuccFrom<N m<n
pigeonhole {succ zero} {succ m} m<n {f} record { property = property } | le x ()
pigeonhole {succ (succ n)} {succ zero} m<n {f} record { property = property } = fzeroNotFsucc (property (transitivity f0 (equalityCommutative f1)))
where
f0 : f (fzero) fzero
f0 with inspect (f fzero)
... | fzero with x = x
... | fsucc () with _
f1 : f (fsucc fzero) fzero
f1 with inspect (f (fsucc fzero))
... | fzero with x = x
... | fsucc () with _
pigeonhole {succ (succ n)} {succ (succ m)} m<n {f} injF = intoSmallerNotSurj {_}{_} {m<n} surj
where
inj : Injection ((intoSmaller m<n) f)
inj = injComp injF (intoSmallerInj m<n)
bij : Bijection ((intoSmaller m<n) f)
bij = injectionIsBijectionFinset inj
surj : Surjection (intoSmaller m<n)
surj = compSurjLeftSurj (Bijection.surj bij)
--surjectionIsBijectionFinset : {n : } {f : FinSet n FinSet n} Surjection f Bijection f
--surjectionIsBijectionFinset {zero} {f} surj = nopointBij f
--surjectionIsBijectionFinset {succ zero} {f} surj = onepointBij f
--surjectionIsBijectionFinset {succ (succ n)} {f} record { property = property } = {!!}
record FiniteSet {a : _} (A : Set a) : Set a where
field
size :
mapping : FinSet size A
bij : Bijection mapping
record InfiniteSet {a : _} (A : Set a) : Set a where
field
isInfinite : (n : ) (f : FinSet n A) (Bijection f) False
finsetInjectInto : {n : } Injection (toNat {n})
Injection.property (finsetInjectInto {zero}) {()}
finsetInjectInto {succ n} = record { property = ans }
where
ans : {n : } {x y : FinSet (succ n)} (toNat x toNat y) x y
ans {zero} {fzero} {fzero} _ = refl
ans {zero} {fzero} {fsucc ()}
ans {zero} {fsucc ()} {y}
ans {succ n} {fzero} {fzero} pr = refl
ans {succ n} {fzero} {fsucc y} ()
ans {succ n} {fsucc x} {fzero} ()
ans {succ n} {fsucc x} {fsucc y} pr with succInjective pr
... | pr' = applyEquality fsucc (ans {n} {x} {y} pr')
IsInfinite : InfiniteSet
InfiniteSet.isInfinite IsInfinite n f bij = pigeonhole (le 0 refl) badInj
where
inv : FinSet n
inv = Invertible.inverse (bijectionImpliesInvertible bij)
invInj : Injection inv
invInj = Bijection.inj (invertibleImpliesBijection (inverseIsInvertible (bijectionImpliesInvertible bij)))
bumpUp : FinSet n FinSet (succ n)
bumpUp = intoSmaller (le 0 refl)
bumpUpInj : Injection bumpUp
bumpUpInj = intoSmallerInj (le 0 refl)
nextInj : Injection (toNat {succ n})
nextInj = finsetInjectInto {succ n}
bad : FinSet (succ n) FinSet n
bad a = (inv (toNat a))
badInj : Injection bad
badInj = injComp nextInj invInj
finsetNotInfinite : {n : } InfiniteSet (FinSet n) False
finsetNotInfinite {n} record { isInfinite = isInfinite } = isInfinite n id idIsBijective

32
FreeGroups.agda Normal file
View File

@@ -0,0 +1,32 @@
{-# OPTIONS --safe --warning=error #-}
open import LogicalFormulae
open import Setoids
open import Functions
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Naturals
open import FinSet
open import Groups
module FreeGroups where
data FreeCompletion {a : _} (A : Set a) : Set a where
letter : A FreeCompletion A
inv : A FreeCompletion A
data FreeGroup {a : _} (A : Set a) : Set a where
emptyWord : FreeGroup A
append : FreeCompletion A FreeGroup A FreeGroup A
reduceWord : {a : _} {A : Set a} (FreeGroup {a} A) FreeGroup A
reduceWord emptyWord = emptyWord
reduceWord (append (letter x) emptyWord) = append (letter x) emptyWord
reduceWord (append (letter x) (append (letter x1) w)) = append (letter x) (reduceWord (append (letter x1) w))
reduceWord (append (letter x) (append (inv x₁) w)) = {!!}
reduceWord (append (inv x) w) = {!!}
freeGroupSetoid : {a : _} (A : Set a) Setoid A
Setoid.__ (freeGroupSetoid A) = {!!}
Setoid.eq (freeGroupSetoid A) = {!!}
freeGroup : {a : _} (A : Set a) Group (freeGroupSetoid A) {!!}
freeGroup A = {!!}

119
Functions.agda Normal file
View File

@@ -0,0 +1,119 @@
{-# OPTIONS --safe --warning=error #-}
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import LogicalFormulae
module Functions where
Rel : {a b : _} Set a Set (a lsuc b)
Rel {a} {b} A = A A Set b
_∘_ : {a b c : _} {A : Set a} {B : Set b} {C : Set c} (f : B C) (g : A B) (A C)
g f = λ a g (f a)
record Injection {a b : _} {A : Set a} {B : Set b} (f : A B) : Set (a b) where
field
property : {x y : A} (f x f y) x y
record Surjection {a b : _} {A : Set a} {B : Set b} (f : A B) : Set (a b) where
field
property : (b : B) Sg A (λ a f a b)
record Bijection {a b : _} {A : Set a} {B : Set b} (f : A B) : Set (a b) where
field
inj : Injection f
surj : Surjection f
record Invertible {a b : _} {A : Set a} {B : Set b} (f : A B) : Set (a b) where
field
inverse : B A
isLeft : (b : B) f (inverse b) b
isRight : (a : A) inverse (f a) a
invertibleImpliesBijection : {a b : _} {A : Set a} {B : Set b} {f : A B} Invertible f Bijection f
Injection.property (Bijection.inj (invertibleImpliesBijection {a} {b} {A} {B} {f} record { inverse = inverse ; isLeft = isLeft ; isRight = isRight })) {x} {y} fx=fy = ans
where
bl : inverse (f x) inverse (f y)
bl = applyEquality inverse fx=fy
ans : x y
ans rewrite equalityCommutative (isRight x) | equalityCommutative (isRight y) = bl
Surjection.property (Bijection.surj (invertibleImpliesBijection {a} {b} {A} {B} {f} record { inverse = inverse ; isLeft = isLeft ; isRight = isRight })) y = (inverse y , isLeft y)
bijectionImpliesInvertible : {a b : _} {A : Set a} {B : Set b} {f : A B} Bijection f Invertible f
Invertible.inverse (bijectionImpliesInvertible record { inj = inj ; surj = record { property = property } }) b = underlying (property b)
Invertible.isLeft (bijectionImpliesInvertible {f = f} record { inj = inj ; surj = surj }) b with Surjection.property surj b
Invertible.isLeft (bijectionImpliesInvertible {f = f} record { inj = inj ; surj = surj }) b | a , prop = prop
Invertible.isRight (bijectionImpliesInvertible {f = f} record { inj = inj ; surj = surj }) a with Surjection.property surj (f a)
Invertible.isRight (bijectionImpliesInvertible {f = f} record { inj = record { property = property } ; surj = surj }) a | a₁ , b = property b
injComp : {a b c : _} {A : Set a} {B : Set b} {C : Set c} {f : A B} {g : B C} Injection f Injection g Injection (g f)
Injection.property (injComp {f = f} {g} record { property = propF } record { property = propG }) pr = propF (propG pr)
surjComp : {a b c : _} {A : Set a} {B : Set b} {C : Set c} {f : A B} {g : B C} Surjection f Surjection g Surjection (g f)
Surjection.property (surjComp {f = f} {g} record { property = propF } record { property = propG }) c with propG c
Surjection.property (surjComp {f = f} {g} record { property = propF } record { property = propG }) c | b , pr with propF b
Surjection.property (surjComp {f = f} {g} record { property = propF } record { property = propG }) c | b , pr | a , pr2 = a , pr'
where
pr' : g (f a) c
pr' rewrite pr2 = pr
bijectionComp : {a b c : _} {A : Set a} {B : Set b} {C : Set c} {f : A B} {g : B C} Bijection f Bijection g Bijection (g f)
Bijection.inj (bijectionComp bijF bijG) = injComp (Bijection.inj bijF) (Bijection.inj bijG)
Bijection.surj (bijectionComp bijF bijG) = surjComp (Bijection.surj bijF) (Bijection.surj bijG)
compInjRightInj : {a b c : _} {A : Set a} {B : Set b} {C : Set c} {f : A B} {g : B C} Injection (g f) Injection f
Injection.property (compInjRightInj {f = f} {g} record { property = property }) {x} {y} fx=fy = property (applyEquality g fx=fy)
compSurjLeftSurj : {a b c : _} {A : Set a} {B : Set b} {C : Set c} {f : A B} {g : B C} Surjection (g f) Surjection g
Surjection.property (compSurjLeftSurj {f = f} {g} record { property = property }) b with property b
Surjection.property (compSurjLeftSurj {f = f} {g} record { property = property }) b | a , b1 = f a , b1
injectionPreservedUnderExtensionalEq : {a b : _} {A : Set a} {B : Set b} {f g : A B} Injection f ({x : A} f x g x) Injection g
Injection.property (injectionPreservedUnderExtensionalEq {A = A} {B} {f} {g} record { property = property } prop) {x} {y} gx=gy = property (transitivity (prop {x}) (transitivity gx=gy (equalityCommutative (prop {y}))))
surjectionPreservedUnderExtensionalEq : {a b : _} {A : Set a} {B : Set b} {f g : A B} Surjection f ({x : A} f x g x) Surjection g
Surjection.property (surjectionPreservedUnderExtensionalEq {f = f} {g} surj ext) b with (Surjection.property surj b)
Surjection.property (surjectionPreservedUnderExtensionalEq {f = f} {g} surj ext) b | a , pr = a , transitivity (equalityCommutative ext) pr
bijectionPreservedUnderExtensionalEq : {a b : _} {A : Set a} {B : Set b} {f g : A B} Bijection f ({x : A} f x g x) Bijection g
Bijection.inj (bijectionPreservedUnderExtensionalEq record { inj = inj ; surj = surj } ext) = injectionPreservedUnderExtensionalEq inj ext
Bijection.surj (bijectionPreservedUnderExtensionalEq record { inj = inj ; surj = surj } ext) = surjectionPreservedUnderExtensionalEq surj ext
inverseIsInvertible : {a b : _} {A : Set a} {B : Set b} {f : A B} (inv : Invertible f) Invertible (Invertible.inverse inv)
Invertible.inverse (inverseIsInvertible {f = f} inv) = f
Invertible.isLeft (inverseIsInvertible {f = f} inv) b = Invertible.isRight inv b
Invertible.isRight (inverseIsInvertible {f = f} inv) a = Invertible.isLeft inv a
id : {a : _} {A : Set a} (A A)
id a = a
idIsBijective : {a : _} {A : Set a} Bijection (id {a} {A})
Injection.property (Bijection.inj idIsBijective) pr = pr
Surjection.property (Bijection.surj idIsBijective) b = b , refl
functionCompositionExtensionallyAssociative : {a b c d : _} {A : Set a} {B : Set b} {C : Set c} {D : Set d} (f : A B) (g : B C) (h : C D) (x : A) (h (g f)) x ((h g) f) x
functionCompositionExtensionallyAssociative f g h x = refl
dom : {a b : _} {A : Set a} {B : Set b} (f : A B) Set a
dom {A = A} f = A
codom : {a b : _} {A : Set a} {B : Set b} (f : A B) Set b
codom {B = B} f = B
record Reflexive {a b : _} {A : Set a} (r : Rel {a} {b} A) : Set (a lsuc b) where
field
reflexive : {b : A} r b b
record Symmetric {a b : _} {A : Set a} (r : Rel {a} {b} A) : Set (a lsuc b) where
field
symmetric : {b c : A} r b c r c b
record Transitive {a b : _} {A : Set a} (r : Rel {a} {b} A) : Set (a lsuc b) where
field
transitive : {b c d : A} r b c r c d r b d
record Equivalence {a b : _} {A : Set a} (r : Rel {a} {b} A) : Set (a lsuc b) where
field
reflexiveEq : Reflexive r
symmetricEq : Symmetric r
transitiveEq : Transitive r

146
GroupActions.agda Normal file
View File

@@ -0,0 +1,146 @@
{-# OPTIONS --safe --warning=error #-}
open import LogicalFormulae
open import Setoids
open import Functions
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Naturals
open import FinSet
open import Groups
module GroupActions where
record GroupAction {m n o p : _} {A : Set m} {S : Setoid {m} {o} A} {_·_ : A A A} {B : Set n} (G : Group S _·_) (X : Setoid {n} {p} B) : Set (m n o p) where
open Group G
open Setoid S renaming (__ to _G_)
open Setoid X renaming (__ to _X_)
field
action : A B B
actionWellDefined1 : {g h : A} {x : B} (g G h) action g x X action h x
actionWellDefined2 : {g : A} {x y : B} (x X y) action g x X action g y
identityAction : {x : B} action identity x X x
associativeAction : {x : B} {g h : A} action (g · h) x X action g (action h x)
trivialAction : {m n o p : _} {A : Set m} {S : Setoid {m} {o} A} {_·_ : A A A} {B : Set n} (G : Group S _·_) (X : Setoid {n} {p} B) GroupAction G X
trivialAction G X = record { action = λ _ x x ; actionWellDefined1 = λ _ reflexive ; actionWellDefined2 = λ wd1 wd1 ; identityAction = reflexive ; associativeAction = reflexive }
where
open Setoid X
open Equivalence eq
open Reflexive reflexiveEq
leftRegularAction : {m n : _} {A : Set m} {S : Setoid {m} {n} A} {_·_ : A A A} (G : Group S _·_) GroupAction G S
GroupAction.action (leftRegularAction {_·_ = _·_} G) g h = g · h
where
open Group G
GroupAction.actionWellDefined1 (leftRegularAction {S = S} G) eq1 = wellDefined eq1 reflexive
where
open Group G
open Setoid S
open Equivalence eq
open Reflexive reflexiveEq
GroupAction.actionWellDefined2 (leftRegularAction {S = S} G) {g} {x} {y} eq1 = wellDefined reflexive eq1
where
open Group G
open Setoid S
open Equivalence eq
open Reflexive reflexiveEq
GroupAction.identityAction (leftRegularAction G) = multIdentLeft
where
open Group G
GroupAction.associativeAction (leftRegularAction {S = S} G) = symmetric multAssoc
where
open Group G
open Setoid S
open Equivalence eq
open Symmetric symmetricEq
conjugationAction : {m n : _} {A : Set m} {S : Setoid {m} {n} A} {_·_ : A A A} (G : Group S _·_) GroupAction G S
conjugationAction {S = S} {_·_ = _·_} G = record { action = λ g h (g · h) · (inverse g) ; actionWellDefined1 = λ gh wellDefined (wellDefined gh reflexive) (inverseWellDefined G gh) ; actionWellDefined2 = λ x~y wellDefined (wellDefined reflexive x~y) reflexive ; identityAction = transitive (wellDefined reflexive (invIdentity G)) (transitive multIdentRight multIdentLeft) ; associativeAction = λ {x} {g} {h} transitive (wellDefined reflexive (invContravariant G g h)) (transitive multAssoc (wellDefined (fourWayAssoc' G) reflexive)) }
where
open Group G
open Setoid S
open Equivalence eq
open Reflexive reflexiveEq
open Transitive transitiveEq
conjugationNormalSubgroupAction : {m n o p : _} {A : Set m} {B : Set o} {S : Setoid {m} {n} A} {T : Setoid {o} {p} B} {_·A_ : A A A} {_·B_ : B B B} (G : Group S _·A_) (H : Group T _·B_) {underF : A B} (f : GroupHom G H underF) GroupAction G (quotientGroupSetoid G f)
GroupAction.action (conjugationNormalSubgroupAction {_·A_ = _·A_} G H {f} fHom) a b = a ·A (b ·A (Group.inverse G a))
GroupAction.actionWellDefined1 (conjugationNormalSubgroupAction {S = S} {T = T} {_·A_ = _·A_} G H {f} fHom) {g} {h} {x} g~h = ans
where
open Group G
open Setoid T
open Equivalence (Setoid.eq T)
open Transitive transitiveEq
open Symmetric (Equivalence.symmetricEq (Setoid.eq S))
open Reflexive (Equivalence.reflexiveEq (Setoid.eq S))
ans : f ((g ·A (x ·A (inverse g))) ·A inverse (h ·A (x ·A inverse h))) Group.identity H
ans = transitive (GroupHom.wellDefined fHom (transferToRight'' G (Group.wellDefined G g~h (Group.wellDefined G reflexive (inverseWellDefined G g~h))))) (imageOfIdentityIsIdentity fHom)
GroupAction.actionWellDefined2 (conjugationNormalSubgroupAction {S = S} {T = T} {_·A_ = _·A_} {_·B_ = _·B_} G H {f} fHom) {g} {x} {y} x~y = ans
where
open Group G
open Setoid T
open Equivalence (Setoid.eq T)
open Transitive transitiveEq
open Transitive (Equivalence.transitiveEq (Setoid.eq S)) renaming (transitive to transitiveG)
open Symmetric (Equivalence.symmetricEq (Setoid.eq S))
open Symmetric (Equivalence.symmetricEq (Setoid.eq T)) renaming (symmetric to symmetricH)
open Reflexive (Equivalence.reflexiveEq (Setoid.eq S))
open Reflexive (Equivalence.reflexiveEq (Setoid.eq T)) renaming (reflexive to reflexiveH)
input : f (x ·A inverse y) Group.identity H
input = x~y
p1 : Setoid.__ S ((g ·A (x ·A inverse g)) ·A inverse (g ·A (y ·A inverse g))) ((g ·A (x ·A inverse g)) ·A (inverse (y ·A (inverse g)) ·A inverse g))
p1 = Group.wellDefined G reflexive (invContravariant G g (y ·A inverse g))
p2 : Setoid.__ S ((g ·A (x ·A inverse g)) ·A (inverse (y ·A (inverse g)) ·A inverse g)) ((g ·A (x ·A inverse g)) ·A ((inverse (inverse g) ·A inverse y) ·A inverse g))
p2 = Group.wellDefined G reflexive (Group.wellDefined G (invContravariant G _ _) reflexive)
p3 : Setoid.__ S ((g ·A (x ·A inverse g)) ·A ((inverse (inverse g) ·A inverse y) ·A inverse g)) (g ·A (((x ·A inverse g) ·A (inverse (inverse g) ·A inverse y)) ·A inverse g))
p3 = symmetric (fourWayAssoc G)
p4 : Setoid.__ S (g ·A (((x ·A inverse g) ·A (inverse (inverse g) ·A inverse y)) ·A inverse g)) (g ·A ((x ·A ((inverse g ·A inverse (inverse g)) ·A inverse y)) ·A inverse g))
p4 = Group.wellDefined G reflexive (Group.wellDefined G (symmetric (fourWayAssoc G)) reflexive)
p5 : Setoid.__ S (g ·A ((x ·A ((inverse g ·A inverse (inverse g)) ·A inverse y)) ·A inverse g)) (g ·A ((x ·A (identity ·A inverse y)) ·A inverse g))
p5 = Group.wellDefined G reflexive (Group.wellDefined G (Group.wellDefined G reflexive (Group.wellDefined G invRight reflexive)) reflexive)
p6 : Setoid.__ S (g ·A ((x ·A (identity ·A inverse y)) ·A inverse g)) (g ·A ((x ·A inverse y) ·A inverse g))
p6 = Group.wellDefined G reflexive (Group.wellDefined G (Group.wellDefined G reflexive multIdentLeft) reflexive)
intermediate : Setoid.__ S ((g ·A (x ·A inverse g)) ·A inverse (g ·A (y ·A inverse g))) (g ·A ((x ·A inverse y) ·A inverse g))
intermediate = transitiveG p1 (transitiveG p2 (transitiveG p3 (transitiveG p4 (transitiveG p5 p6))))
p7 : f ((g ·A (x ·A inverse g)) ·A inverse (g ·A (y ·A inverse g))) f (g ·A ((x ·A inverse y) ·A inverse g))
p7 = GroupHom.wellDefined fHom intermediate
p8 : f (g ·A ((x ·A inverse y) ·A inverse g)) (f g) ·B (f ((x ·A inverse y) ·A inverse g))
p8 = GroupHom.groupHom fHom
p9 : (f g) ·B (f ((x ·A inverse y) ·A inverse g)) (f g) ·B (f (x ·A inverse y) ·B f (inverse g))
p9 = Group.wellDefined H reflexiveH (GroupHom.groupHom fHom)
p10 : (f g) ·B (f (x ·A inverse y) ·B f (inverse g)) (f g) ·B (Group.identity H ·B f (inverse g))
p10 = Group.wellDefined H reflexiveH (Group.wellDefined H input reflexiveH)
p11 : (f g) ·B (Group.identity H ·B f (inverse g)) (f g) ·B (f (inverse g))
p11 = Group.wellDefined H reflexiveH (Group.multIdentLeft H)
p12 : (f g) ·B (f (inverse g)) f (g ·A (inverse g))
p12 = symmetricH (GroupHom.groupHom fHom)
intermediate2 : f ((g ·A (x ·A inverse g)) ·A inverse (g ·A (y ·A inverse g))) (f (g ·A (inverse g)))
intermediate2 = transitive p7 (transitive p8 (transitive p9 (transitive p10 (transitive p11 p12))))
ans : f ((g ·A (x ·A inverse g)) ·A inverse (g ·A (y ·A inverse g))) Group.identity H
ans = transitive intermediate2 (transitive (GroupHom.wellDefined fHom invRight) (imageOfIdentityIsIdentity fHom))
GroupAction.identityAction (conjugationNormalSubgroupAction {S = S} {T = T} {_·A_ = _·A_} G H {f} fHom) {x} = ans
where
open Group G
open Setoid S
open Setoid T renaming (__ to _T_)
open Equivalence (Setoid.eq T)
open Transitive transitiveEq
i : Setoid.__ S (x ·A inverse identity) x
i = Transitive.transitive (Equivalence.transitiveEq (Setoid.eq S)) (wellDefined (Reflexive.reflexive (Equivalence.reflexiveEq (Setoid.eq S))) (invIdentity G)) multIdentRight
h : identity ·A (x ·A inverse identity) x
h = Transitive.transitive (Equivalence.transitiveEq (Setoid.eq S)) multIdentLeft i
g : ((identity ·A (x ·A inverse identity)) ·A inverse x) identity
g = transferToRight'' G h
ans : f ((identity ·A (x ·A inverse identity)) ·A Group.inverse G x) T Group.identity H
ans = transitive (GroupHom.wellDefined fHom g) (imageOfIdentityIsIdentity fHom)
GroupAction.associativeAction (conjugationNormalSubgroupAction {S = S} {T = T} {_·A_ = _·A_} G H {f} fHom) {x} {g} {h} = ans
where
open Group G
open Setoid T renaming (__ to _T_)
open Setoid S renaming (__ to _S_)
open Transitive (Equivalence.transitiveEq (Setoid.eq T)) renaming (transitive to transitiveH)
open Transitive (Equivalence.transitiveEq (Setoid.eq S)) renaming (transitive to transitiveG)
open Symmetric (Equivalence.symmetricEq (Setoid.eq S)) renaming (symmetric to symmetricG)
open Reflexive (Equivalence.reflexiveEq (Setoid.eq S))
ans : f (((g ·A h) ·A (x ·A inverse (g ·A h))) ·A inverse ((g ·A ((h ·A (x ·A inverse h)) ·A inverse g)))) T Group.identity H
ans = transitiveH (GroupHom.wellDefined fHom (transferToRight'' G (transitiveG (symmetricG multAssoc) (Group.wellDefined G reflexive (transitiveG (wellDefined reflexive (transitiveG (wellDefined reflexive (invContravariant G g h)) multAssoc)) multAssoc))))) (imageOfIdentityIsIdentity fHom)

50
GroupCosets.agda Normal file
View File

@@ -0,0 +1,50 @@
{-# OPTIONS --safe --warning=error #-}
open import LogicalFormulae
open import Setoids
open import Functions
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Naturals
open import Integers
open import FinSet
open import Groups
module GroupCosets where
data GroupCoset {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+A_ : A A A} {_+B_ : B B B} {G : Group S _+A_} {H : Group T _+B_} {f : B A} {fHom : GroupHom H G f} (subgrp : Subgroup G H fHom) : Set (a b c d) where
cosetOfElt : (x : A) GroupCoset subgrp
groupCosetSetoid : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+A_ : A A A} {_+B_ : B B B} {G : Group S _+A_} {H : Group T _+B_} {f : B A} {fHom : GroupHom H G f} (subgrp : Subgroup G H fHom) Setoid (GroupCoset subgrp)
Setoid.__ (groupCosetSetoid {A = A} {S = S} {_+A_ = _+A_} {G = G} subgrp) (cosetOfElt x) (cosetOfElt y) = Sg (A && A) (λ p x +A (_&&_.fst p) y +A (_&&_.snd p))
where
open Setoid S
Reflexive.reflexive (Equivalence.reflexiveEq (Setoid.eq (groupCosetSetoid {S = S} {G = G} subgrp))) {cosetOfElt x} = (identity ,, identity) , transitive (multIdentRight) (symmetric multIdentRight)
where
open Group G
open Setoid S
open Transitive (Equivalence.transitiveEq eq)
open Symmetric (Equivalence.symmetricEq eq)
Symmetric.symmetric (Equivalence.symmetricEq (Setoid.eq (groupCosetSetoid {S = S} subgrp))) {cosetOfElt x} {cosetOfElt y} ((xMul ,, yMul) , pr) = (yMul ,, xMul) , Symmetric.symmetric (Equivalence.symmetricEq (Setoid.eq S)) pr
Transitive.transitive (Equivalence.transitiveEq (Setoid.eq (groupCosetSetoid {S = S} {_+A_ = _+A_} {G = G} subgrp))) {cosetOfElt x} {cosetOfElt y} {cosetOfElt z} ((xMul ,, yMul) , pr) ((yMul' ,, zMul) , pr2) = (((xMul +A (inverse yMul)) +A yMul') ,, zMul) , transitive (transitive multAssoc (wellDefined (transitive multAssoc (transitive (wellDefined pr reflexive) (transitive (symmetric multAssoc) (transitive (wellDefined reflexive invRight) multIdentRight)))) reflexive)) pr2
where
open Group G
open Setoid S
open Transitive (Equivalence.transitiveEq eq)
open Reflexive (Equivalence.reflexiveEq eq)
open Symmetric (Equivalence.symmetricEq eq)
groupCosetBijection : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+A_ : A A A} {_+B_ : B B B} {G : Group S _+A_} {H : Group T _+B_} {f : B A} {fHom : GroupHom H G f} (subgrp : Subgroup G H fHom) SetoidsBiject T (groupCosetSetoid subgrp)
SetoidsBiject.bij (groupCosetBijection {f = f} subgrp) x = cosetOfElt (f x)
SetoidInjection.wellDefined (SetoidBijection.inj (SetoidsBiject.bijIsBijective (groupCosetBijection {S = S} {G = G} {fHom = fHom} subgrp))) x~y = (identity ,, identity) , transitive multIdentRight (transitive (GroupHom.wellDefined fHom x~y) (symmetric multIdentRight))
where
open Group G
open Setoid S
open Transitive (Equivalence.transitiveEq eq)
open Symmetric (Equivalence.symmetricEq eq)
SetoidInjection.injective (SetoidBijection.inj (SetoidsBiject.bijIsBijective (groupCosetBijection {S = S} {_+A_ = _+A_} {G = G} {f = f} subgrp))) {x} {y} ((xH ,, yH) , b) = {!!}
SetoidSurjection.wellDefined (SetoidBijection.surj (SetoidsBiject.bijIsBijective (groupCosetBijection {S = S} {G = G} {fHom = fHom} subgrp))) x~y = (identity ,, identity) , transitive multIdentRight (transitive (GroupHom.wellDefined fHom x~y) (symmetric multIdentRight))
where
open Group G
open Setoid S
open Transitive (Equivalence.transitiveEq eq)
open Symmetric (Equivalence.symmetricEq eq)
SetoidSurjection.surjective (SetoidBijection.surj (SetoidsBiject.bijIsBijective (groupCosetBijection subgrp))) {cosetOfElt x} = {!!}

488
Groups.agda Normal file
View File

@@ -0,0 +1,488 @@
{-# OPTIONS --safe --warning=error #-}
open import LogicalFormulae
open import Setoids
open import Functions
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Naturals
open import FinSet
module Groups where
record Group {lvl1 lvl2} {A : Set lvl1} (S : Setoid {lvl1} {lvl2} A) (_·_ : A A A) : Set (lsuc lvl1 lvl2) where
open Setoid S
field
wellDefined : {m n x y : A} (m x) (n y) (m · n) (x · y)
identity : A
inverse : A A
multAssoc : {a b c : A} (a · (b · c)) (a · b) · c
multIdentRight : {a : A} (a · identity) a
multIdentLeft : {a : A} (identity · a) a
invLeft : {a : A} (inverse a) · a identity
invRight : {a : A} a · (inverse a) identity
reflGroupWellDefined : {lvl : _} {A : Set lvl} {m n x y : A} {op : A A A} m x n y (op m n) (op x y)
reflGroupWellDefined {lvl} {A} {m} {n} {.m} {.n} {op} refl refl = refl
record AbelianGroup {a} {b} {A : Set a} {S : Setoid {a} {b} A} {_·_ : A A A} (G : Group S _·_) : Set (lsuc a b) where
open Setoid S
field
commutative : {a b : A} (a · b) (b · a)
groupsHaveLeftCancellation : {a b : _} {A : Set a} {_·_ : A A A} {S : Setoid {a} {b} A} (G : Group S _·_) (x y z : A) (Setoid.__ S (x · y) (x · z)) (Setoid.__ S y z)
groupsHaveLeftCancellation {_·_ = _·_} {S} g x y z pr = o
where
open Group g renaming (inverse to _^-1) renaming (identity to e)
open Setoid S
open Transitive (Equivalence.transitiveEq eq)
open Symmetric (Equivalence.symmetricEq eq)
j : ((x ^-1) · x) · y (x ^-1) · (x · z)
j = transitive (symmetric (multAssoc {x ^-1} {x} {y})) (wellDefined ~refl pr)
k : ((x ^-1) · x) · y ((x ^-1) · x) · z
k = transitive j multAssoc
l : e · y ((x ^-1) · x) · z
l = transitive (wellDefined (symmetric invLeft) ~refl) k
m : e · y e · z
m = transitive l (wellDefined invLeft ~refl)
n : y e · z
n = transitive (symmetric multIdentLeft) m
o : y z
o = transitive n multIdentLeft
groupsHaveRightCancellation : {a b : _} {A : Set a} {_·_ : A A A} {S : Setoid {a} {b} A} (G : Group S _·_) (x y z : A) (Setoid.__ S (y · x) (z · x)) (Setoid.__ S y z)
groupsHaveRightCancellation {_·_ = _·_} {S} g x y z pr = transitive m multIdentRight
where
open Group g renaming (inverse to _^-1) renaming (identity to e)
open Setoid S
open Transitive (Equivalence.transitiveEq eq)
open Symmetric (Equivalence.symmetricEq eq)
i : (y · x) · (x ^-1) (z · x) · (x ^-1)
i = wellDefined pr ~refl
j : y · (x · (x ^-1)) (z · x) · (x ^-1)
j = transitive multAssoc i
j' : y · e (z · x) · (x ^-1)
j' = transitive (wellDefined ~refl (symmetric invRight)) j
k : y (z · x) · (x ^-1)
k = transitive (symmetric multIdentRight) j'
l : y z · (x · (x ^-1))
l = transitive k (symmetric multAssoc)
m : y z · e
m = transitive l (wellDefined ~refl invRight)
identityIsUnique : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_·_ : A A A} (G : Group S _·_) (e : A) ((b : A) (Setoid.__ S (b · e) b)) (Setoid.__ S e (Group.identity G))
identityIsUnique {S = S} {_·_} g thing fb = transitive (symmetric multIdentLeft) (fb e)
where
open Group g renaming (inverse to _^-1) renaming (identity to e)
open Setoid S
open Transitive (Equivalence.transitiveEq eq)
open Symmetric (Equivalence.symmetricEq eq)
replaceGroupOp : {l m : _} {A : Set l} {S : Setoid {l} {m} A} {_·_ : A A A} (G : Group S _·_) {a b c d w x y z : A} (Setoid.__ S a c) (Setoid.__ S b d) (Setoid.__ S w y) (Setoid.__ S x z) Setoid.__ S (a · b) (w · x) Setoid.__ S (c · d) (y · z)
replaceGroupOp {S = S} {_·_} G a~c b~d w~y x~z pr = transitive (symmetric (wellDefined a~c b~d)) (transitive pr (wellDefined w~y x~z))
where
open Group G
open Setoid S
open Equivalence eq
open Transitive transitiveEq
open Symmetric symmetricEq
replaceGroupOpRight : {l m : _} {A : Set l} {S : Setoid {l} {m} A} {_·_ : A A A} (G : Group S _·_) {a b c x : A} (Setoid.__ S a (b · c)) (Setoid.__ S c x) (Setoid.__ S a (b · x))
replaceGroupOpRight {S = S} {_·_} G a~bc c~x = transitive a~bc (wellDefined reflexive c~x)
where
open Group G
open Setoid S
open Equivalence eq
open Transitive transitiveEq
open Reflexive reflexiveEq
inverseWellDefined : {l m : _} {A : Set l} {S : Setoid {l} {m} A} {_·_ : A A A} (G : Group S _·_) {x y : A} (Setoid.__ S x y) Setoid.__ S (Group.inverse G x) (Group.inverse G y)
inverseWellDefined {S = S} {_·_} G {x} {y} x~y = groupsHaveRightCancellation G x (inverse x) (inverse y) q
where
open Group G
open Setoid S
open Equivalence eq
open Transitive transitiveEq
open Symmetric symmetricEq
p : inverse x · x inverse y · y
p = transitive invLeft (symmetric invLeft)
q : inverse x · x inverse y · x
q = replaceGroupOpRight G {_·_ (inverse x) x} {inverse y} {y} {x} p (symmetric x~y)
rightInversesAreUnique : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_·_ : A A A} (G : Group S _·_) (x : A) (y : A) Setoid.__ S (y · x) (Group.identity G) Setoid.__ S y (Group.inverse G x)
rightInversesAreUnique {S = S} {_·_} G x y f = transitive i (transitive j (transitive k (transitive l m)))
where
open Group G renaming (inverse to _^-1) renaming (identity to e)
open Setoid S
open Transitive (Equivalence.transitiveEq eq)
open Symmetric (Equivalence.symmetricEq eq)
i : y y · e
j : y · e y · (x · (x ^-1))
k : y · (x · (x ^-1)) (y · x) · (x ^-1)
l : (y · x) · (x ^-1) e · (x ^-1)
m : e · (x ^-1) x ^-1
i = symmetric multIdentRight
j = wellDefined ~refl (symmetric invRight)
k = multAssoc
l = wellDefined f ~refl
m = multIdentLeft
leftInversesAreUnique : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_·_ : A A A} (G : Group S _·_) {x : A} {y : A} Setoid.__ S (x · y) (Group.identity G) Setoid.__ S y (Group.inverse G x)
leftInversesAreUnique {S = S} {_·_} G {x} {y} f = rightInversesAreUnique G x y l
where
open Group G renaming (inverse to _^-1) renaming (identity to e)
open Setoid S
open Transitive (Equivalence.transitiveEq eq)
open Symmetric (Equivalence.symmetricEq eq)
i : y · (x · y) y · e
i' : y · (x · y) y
j : (y · x) · y y
k : (y · x) · y e · y
l : y · x e
i = wellDefined ~refl f
i' = transitive i multIdentRight
j = transitive (symmetric multAssoc) i'
k = transitive j (symmetric multIdentLeft)
l = groupsHaveRightCancellation G y (y · x) e k
invTwice : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_·_ : A A A} (G : Group S _·_) (x : A) Setoid.__ S (Group.inverse G (Group.inverse G x)) x
invTwice {S = S} {_·_} G x = symmetric (rightInversesAreUnique G (x ^-1) x invRight)
where
open Group G renaming (inverse to _^-1) renaming (identity to e)
open Setoid S
open Transitive (Equivalence.transitiveEq eq)
open Symmetric (Equivalence.symmetricEq eq)
fourWayAssoc : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_·_ : A A A} (G : Group S _·_) {r s t u : A} (Setoid.__ S) (r · ((s · t) · u)) ((r · s) · (t · u))
fourWayAssoc {S = S} {_·_} G {r} {s} {t} {u} = transitive p1 (transitive p2 p3)
where
open Group G renaming (inverse to _^-1) renaming (identity to e)
open Setoid S
open Equivalence eq
open Transitive transitiveEq
open Reflexive reflexiveEq
open Symmetric symmetricEq
p1 : r · ((s · t) · u) (r · (s · t)) · u
p2 : (r · (s · t)) · u ((r · s) · t) · u
p3 : ((r · s) · t) · u (r · s) · (t · u)
p1 = Group.multAssoc G
p2 = Group.wellDefined G (Group.multAssoc G) reflexive
p3 = symmetric (Group.multAssoc G)
fourWayAssoc' : {m n : _} {A : Set m} {S : Setoid {m} {n} A} {_·_ : A A A} (G : Group S _·_) {a b c d : A} (Setoid.__ S (((a · b) · c) · d) (a · ((b · c) · d)))
fourWayAssoc' {S = S} G = transitive (symmetric multAssoc) (symmetric (fourWayAssoc G))
where
open Group G
open Setoid S
open Equivalence eq
open Transitive transitiveEq
open Symmetric symmetricEq
fourWayAssoc'' : {m n : _} {A : Set m} {S : Setoid {m} {n} A} {_·_ : A A A} (G : Group S _·_) {a b c d : A} (Setoid.__ S (a · (b · (c · d))) (a · ((b · c) · d)))
fourWayAssoc'' {S = S} {_·_ = _·_} G = transitive multAssoc (symmetric (fourWayAssoc G))
where
open Group G
open Setoid S
open Equivalence eq
open Transitive transitiveEq
open Symmetric symmetricEq
invContravariant : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_·_ : A A A} (G : Group S _·_) (x y : A) (Setoid.__ S (Group.inverse G (x · y)) ((Group.inverse G y) · (Group.inverse G x)))
invContravariant {S = S} {_·_} G x y = ans
where
open Group G renaming (inverse to _^-1) renaming (identity to e)
open Setoid S
open Equivalence eq
open Transitive transitiveEq
open Symmetric symmetricEq
open Reflexive reflexiveEq
otherInv = (y ^-1) · (x ^-1)
manyAssocs : x · ((y · (y ^-1)) · (x ^-1)) (x · y) · ((y ^-1) · (x ^-1))
oneMult : (x · y) · otherInv x · (x ^-1)
manyAssocs = fourWayAssoc G
oneMult = symmetric (transitive (Group.wellDefined G reflexive (transitive (symmetric (Group.multIdentLeft G)) (Group.wellDefined G (symmetric (Group.invRight G)) reflexive))) manyAssocs)
otherInvIsInverse : (x · y) · otherInv e
otherInvIsInverse = transitive oneMult (Group.invRight G)
ans : (x · y) ^-1 (y ^-1) · (x ^-1)
ans = symmetric (leftInversesAreUnique G otherInvIsInverse)
invIdentity : {l m : _} {A : Set l} {S : Setoid {l} {m} A} {_·_ : A A A} (G : Group S _·_) Setoid.__ S ((Group.inverse G) (Group.identity G)) (Group.identity G)
invIdentity {S = S} G = transitive (symmetric multIdentLeft) invRight
where
open Group G
open Setoid S
open Equivalence eq
open Symmetric symmetricEq
open Transitive transitiveEq
transferToRight : {m n : _} {A : Set m} {S : Setoid {m} {n} A} {_·_ : A A A} (G : Group S _·_) {a b : A} Setoid.__ S (a · (Group.inverse G b)) (Group.identity G) Setoid.__ S a b
transferToRight {S = S} {_·_ = _·_} G {a} {b} ab-1 = transitive (symmetric (invTwice G a)) (transitive u (invTwice G b))
where
open Setoid S
open Group G
open Equivalence eq
open Symmetric symmetricEq
open Transitive transitiveEq
t : inverse a inverse b
t = symmetric (leftInversesAreUnique G ab-1)
u : inverse (inverse a) inverse (inverse b)
u = inverseWellDefined G t
transferToRight' : {m n : _} {A : Set m} {S : Setoid {m} {n} A} {_·_ : A A A} (G : Group S _·_) {a b : A} Setoid.__ S (a · b) (Group.identity G) Setoid.__ S a (Group.inverse G b)
transferToRight' {S = S} {_·_ = _·_} G {a} {b} ab-1 = transferToRight G lemma
where
open Setoid S
open Group G
open Equivalence eq
open Symmetric symmetricEq
open Transitive transitiveEq
open Reflexive reflexiveEq
lemma : a · (inverse (inverse b)) identity
lemma = transitive (wellDefined reflexive (invTwice G b)) ab-1
transferToRight'' : {m n : _} {A : Set m} {S : Setoid {m} {n} A} {_·_ : A A A} (G : Group S _·_) {a b : A} Setoid.__ S a b Setoid.__ S (a · (Group.inverse G b)) (Group.identity G)
transferToRight'' {S = S} {_·_ = _·_} G {a} {b} a~b = transitive (Group.wellDefined G a~b reflexive) invRight
where
open Group G
open Setoid S
open Transitive (Equivalence.transitiveEq eq)
open Reflexive (Equivalence.reflexiveEq eq)
directSumGroup : {m n o p : _} {A : Set m} {S : Setoid {m} {o} A} {_·A_ : A A A} {B : Set n} {T : Setoid {n} {p} B} {_·B_ : B B B} (G : Group S _·A_) (h : Group T _·B_) Group (directSumSetoid S T) (λ x1 y1 (((_&&_.fst x1) ·A (_&&_.fst y1)) ,, ((_&&_.snd x1) ·B (_&&_.snd y1))))
Group.wellDefined (directSumGroup {A = A} {B} g h) (s ,, t) (u ,, v) = Group.wellDefined g s u ,, Group.wellDefined h t v
Group.identity (directSumGroup {A = A} {B} g h) = (Group.identity g ,, Group.identity h)
Group.inverse (directSumGroup {A = A} {B} g h) (g1 ,, h1) = (Group.inverse g g1) ,, (Group.inverse h h1)
Group.multAssoc (directSumGroup {A = A} {B} g h) = Group.multAssoc g ,, Group.multAssoc h
Group.multIdentRight (directSumGroup {A = A} {B} g h) = Group.multIdentRight g ,, Group.multIdentRight h
Group.multIdentLeft (directSumGroup {A = A} {B} g h) = Group.multIdentLeft g ,, Group.multIdentLeft h
Group.invLeft (directSumGroup {A = A} {B} g h) = Group.invLeft g ,, Group.invLeft h
Group.invRight (directSumGroup {A = A} {B} g h) = Group.invRight g ,, Group.invRight h
directSumAbelianGroup : {m n o p : _} {A : Set m} {S : Setoid {m} {o} A} {_·A_ : A A A} {B : Set n} {T : Setoid {n} {p} B} {_·B_ : B B B} {underG : Group S _·A_} {underH : Group T _·B_} (G : AbelianGroup underG) (h : AbelianGroup underH) (AbelianGroup (directSumGroup underG underH))
AbelianGroup.commutative (directSumAbelianGroup {A = A} {B} G H) = AbelianGroup.commutative G ,, AbelianGroup.commutative H
record GroupHom {m n o p : _} {A : Set m} {S : Setoid {m} {o} A} {_·A_ : A A A} {B : Set n} {T : Setoid {n} {p} B} {_·B_ : B B B} (G : Group S _·A_) (H : Group T _·B_) (f : A B) : Set (m n o p) where
open Group H
open Setoid T
field
groupHom : {x y : A} f (x ·A y) (f x) ·B (f y)
wellDefined : {x y : A} Setoid.__ S x y f x f y
record InjectiveGroupHom {m n o p : _} {A : Set m} {S : Setoid {m} {o} A} {_·A_ : A A A} {B : Set n} {T : Setoid {n} {p} B} {_·B_ : B B B} {G : Group S _·A_} {H : Group T _·B_} {underf : A B} (f : GroupHom G H underf) : Set (m n o p) where
open Setoid S renaming (__ to _A_)
open Setoid T renaming (__ to _B_)
field
injective : SetoidInjection S T underf
imageOfIdentityIsIdentity : {m n o p : _} {A : Set m} {S : Setoid {m} {o} A} {B : Set n} {T : Setoid {n} {p} B} {_·A_ : A A A} {_·B_ : B B B} {G : Group S _·A_} {H : Group T _·B_} {f : A B} (hom : GroupHom G H f) Setoid.__ T (f (Group.identity G)) (Group.identity H)
imageOfIdentityIsIdentity {S = S} {T = T} {_·A_ = _·A_} {_·B_ = _·B_} {G = G} {H = H} {f = f} hom = Symmetric.symmetric (Equivalence.symmetricEq eq) t
where
open Group H
open Setoid T
id2 : Setoid.__ S (Group.identity G) ((Group.identity G) ·A (Group.identity G))
id2 = Symmetric.symmetric (Equivalence.symmetricEq (Setoid.eq S)) (Group.multIdentRight G)
r : f (Group.identity G) f (Group.identity G) ·B f (Group.identity G)
s : identity ·B f (Group.identity G) f (Group.identity G) ·B f (Group.identity G)
t : identity f (Group.identity G)
t = groupsHaveRightCancellation H (f (Group.identity G)) identity (f (Group.identity G)) s
s = Transitive.transitive (Equivalence.transitiveEq eq) multIdentLeft r
r = Transitive.transitive (Equivalence.transitiveEq eq) (GroupHom.wellDefined hom id2) (GroupHom.groupHom hom)
groupHomsCompose : {m n o r s t : _} {A : Set m} {S : Setoid {m} {r} A} {_+A_ : A A A} {B : Set n} {T : Setoid {n} {s} B} {_+B_ : B B B} {C : Set o} {U : Setoid {o} {t} C} {_+C_ : C C C} {G : Group S _+A_} {H : Group T _+B_} {I : Group U _+C_} {f : A B} {g : B C} (fHom : GroupHom G H f) (gHom : GroupHom H I g) GroupHom G I (g f)
GroupHom.wellDefined (groupHomsCompose {G = G} {H} {I} {f} {g} fHom gHom) {x} {y} pr = GroupHom.wellDefined gHom (GroupHom.wellDefined fHom pr)
GroupHom.groupHom (groupHomsCompose {S = S} {_+A_ = _·A_} {T = T} {U = U} {_+C_ = _·C_} {G = G} {H} {I} {f} {g} fHom gHom) {x} {y} = answer
where
open Group I
answer : (Setoid.__ U) ((g f) (x ·A y)) ((g f) x ·C (g f) y)
answer = (Transitive.transitive (Equivalence.transitiveEq (Setoid.eq U))) (GroupHom.wellDefined gHom (GroupHom.groupHom fHom {x} {y}) ) (GroupHom.groupHom gHom {f x} {f y})
record GroupIso {m n o p : _} {A : Set m} {S : Setoid {m} {o} A} {_·A_ : A A A} {B : Set n} {T : Setoid {n} {p} B} {_·B_ : B B B} (G : Group S _·A_) (H : Group T _·B_) (f : A B) : Set (m n o p) where
open Setoid S renaming (__ to _G_)
open Setoid T renaming (__ to _H_)
field
groupHom : GroupHom G H f
bij : SetoidBijection S T f
record GroupsIsomorphic {m n o p : _} {A : Set m} {S : Setoid {m} {o} A} {_·A_ : A A A} {B : Set n} {T : Setoid {n} {p} B} {_·B_ : B B B} (G : Group S _·A_) (H : Group T _·B_) : Set (m n o p) where
field
isomorphism : A B
proof : GroupIso G H isomorphism
groupIsosCompose : {m n o r s t : _} {A : Set m} {S : Setoid {m} {r} A} {_+A_ : A A A} {B : Set n} {T : Setoid {n} {s} B} {_+B_ : B B B} {C : Set o} {U : Setoid {o} {t} C} {_+C_ : C C C} {G : Group S _+A_} {H : Group T _+B_} {I : Group U _+C_} {f : A B} {g : B C} (fHom : GroupIso G H f) (gHom : GroupIso H I g) GroupIso G I (g f)
GroupIso.groupHom (groupIsosCompose fHom gHom) = groupHomsCompose (GroupIso.groupHom fHom) (GroupIso.groupHom gHom)
GroupIso.bij (groupIsosCompose {A = A} {S = S} {T = T} {C = C} {U = U} {f = f} {g = g} fHom gHom) = record { inj = record { injective = λ pr (SetoidInjection.injective (SetoidBijection.inj (GroupIso.bij fHom))) (SetoidInjection.injective (SetoidBijection.inj (GroupIso.bij gHom)) pr) ; wellDefined = wellDefined } ; surj = record { surjective = surj ; wellDefined = wellDefined } }
where
open Setoid S renaming (__ to _A_)
open Setoid U renaming (__ to _C_)
wellDefined : {x y : A} (x A y) (g (f x) C g (f y))
wellDefined = GroupHom.wellDefined (groupHomsCompose (GroupIso.groupHom fHom) (GroupIso.groupHom gHom))
surj : {x : C} Sg A (λ a (g (f a) C x))
surj {x} with SetoidSurjection.surjective (SetoidBijection.surj (GroupIso.bij gHom)) {x}
surj {x} | a , prA with SetoidSurjection.surjective (SetoidBijection.surj (GroupIso.bij fHom)) {a}
... | b , prB = b , transitive (GroupHom.wellDefined (GroupIso.groupHom gHom) prB) prA
where
open Transitive (Equivalence.transitiveEq (Setoid.eq U))
record FiniteGroup {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_·_ : A A A} (G : Group S _·_) {c : _} (quotientSet : Set c) : Set (a b c) where
field
toSet : SetoidToSet S quotientSet
finite : FiniteSet quotientSet
groupOrder : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_·_ : A A A} {underG : Group S _·_} {c : _} {quotientSet : Set c} (G : FiniteGroup underG quotientSet)
groupOrder record { toSet = toSet ; finite = record { size = size ; mapping = mapping ; bij = bij } } = size
--groupIsoInvertible : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d}} {_+A_ : A A A} {_+B_ : B B B} {G : Group S _+A_} {H : Group T _+B_} {f : A B} (iso : GroupIso G H f) GroupIso H G (Invertible.inverse (bijectionImpliesInvertible (GroupIso.bijective iso)))
--GroupHom.groupHom (GroupIso.groupHom (groupIsoInvertible {G = G} {H} {f} iso)) {x} {y} = {!!}
-- where
-- open Group G renaming (_·_ to _+G_)
-- open Group H renaming (_·_ to _+H_)
--GroupHom.wellDefined (GroupIso.groupHom (groupIsoInvertible {G = G} {H} {f} iso)) {x} {y} xy = {!GroupHom.wellDefined xy!}
--GroupIso.bijective (groupIsoInvertible {G = G} {H} {f} iso) = invertibleImpliesBijection (inverseIsInvertible (bijectionImpliesInvertible (GroupIso.bijective iso)))
homRespectsInverse : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_·A_ : A A A} {_·B_ : B B B} {G : Group S _·A_} {H : Group T _·B_} {underF : A B} (f : GroupHom G H underF) {x : A} Setoid.__ T (underF (Group.inverse G x)) (Group.inverse H (underF x))
homRespectsInverse {T = T} {_·A_ = _·A_} {_·B_ = _·B_} {G = G} {H = H} {underF = f} fHom {x} = rightInversesAreUnique H (f x) (f (Group.inverse G x)) (transitive (symmetric (GroupHom.groupHom fHom)) (transitive (GroupHom.wellDefined fHom (Group.invLeft G)) (imageOfIdentityIsIdentity fHom)))
where
open Setoid T
open Equivalence eq
open Transitive transitiveEq
open Symmetric symmetricEq
quotientGroupSetoid : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_·A_ : A A A} {_·B_ : B B B} (G : Group S _·A_) {H : Group T _·B_} {underf : A B} (f : GroupHom G H underf) (Setoid {a} {d} A)
quotientGroupSetoid {A = A} {S = S} {T = T} {_·A_ = _·A_} {_·B_ = _·B_} G {H} {f} fHom = ansSetoid
where
open Setoid T
open Group H
open Equivalence eq
open Transitive transitiveEq
open Reflexive reflexiveEq
ansSetoid : Setoid A
Setoid.__ ansSetoid r s = (f (r ·A (Group.inverse G s))) identity
Reflexive.reflexive (Equivalence.reflexiveEq (Setoid.eq ansSetoid)) {b} = transitive (GroupHom.wellDefined fHom (Group.invRight G)) (imageOfIdentityIsIdentity fHom)
Symmetric.symmetric (Equivalence.symmetricEq (Setoid.eq ansSetoid)) {m} {n} pr = i
where
g : f (Group.inverse G (m ·A Group.inverse G n)) identity
g = transitive (homRespectsInverse fHom {m ·A Group.inverse G n}) (transitive (inverseWellDefined H pr) (invIdentity H))
h : f (Group.inverse G (Group.inverse G n) ·A Group.inverse G m) identity
h = transitive (GroupHom.wellDefined fHom (Symmetric.symmetric (Equivalence.symmetricEq (Setoid.eq S)) (invContravariant G m (Group.inverse G n)))) g
i : f (n ·A Group.inverse G m) identity
i = transitive (GroupHom.wellDefined fHom (Group.wellDefined G (Symmetric.symmetric (Equivalence.symmetricEq (Setoid.eq S)) (invTwice G n)) (Reflexive.reflexive (Equivalence.reflexiveEq (Setoid.eq S))))) h
Transitive.transitive (Equivalence.transitiveEq (Setoid.eq ansSetoid)) {m} {n} {o} prmn prno = transitive (GroupHom.wellDefined fHom (Group.wellDefined G (Reflexive.reflexive (Equivalence.reflexiveEq (Setoid.eq S))) (Symmetric.symmetric (Equivalence.symmetricEq (Setoid.eq S)) (Group.multIdentLeft G)))) k
where
g : f (m ·A Group.inverse G n) ·B f (n ·A Group.inverse G o) identity ·B identity
g = replaceGroupOp H reflexive reflexive prmn prno reflexive
h : f (m ·A Group.inverse G n) ·B f (n ·A Group.inverse G o) identity
h = transitive g multIdentLeft
i : f ((m ·A Group.inverse G n) ·A (n ·A Group.inverse G o)) identity
i = transitive (GroupHom.groupHom fHom) h
j : f (m ·A (((Group.inverse G n) ·A n) ·A Group.inverse G o)) identity
j = transitive (GroupHom.wellDefined fHom (fourWayAssoc G)) i
k : f (m ·A ((Group.identity G) ·A Group.inverse G o)) identity
k = transitive (GroupHom.wellDefined fHom (Group.wellDefined G (Reflexive.reflexive (Equivalence.reflexiveEq (Setoid.eq S))) (Group.wellDefined G (Symmetric.symmetric (Equivalence.symmetricEq (Setoid.eq S)) (Group.invLeft G)) (Reflexive.reflexive (Equivalence.reflexiveEq (Setoid.eq S)))))) j
quotientGroup : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_·A_ : A A A} {_·B_ : B B B} (G : Group S _·A_) {H : Group T _·B_} {underf : A B} (f : GroupHom G H underf) Group (quotientGroupSetoid G f) _·A_
Group.wellDefined (quotientGroup {S = S} {T = T} {_·A_ = _·A_} {_·B_ = _·B_} G {H = H} {underf = f} fHom) {x} {y} {m} {n} x~m y~n = ans
where
open Setoid T
open Equivalence (Setoid.eq T)
open Symmetric symmetricEq
open Transitive transitiveEq
p1 : f ((x ·A y) ·A (Group.inverse G (m ·A n))) f ((x ·A y) ·A ((Group.inverse G n) ·A (Group.inverse G m)))
p2 : f ((x ·A y) ·A ((Group.inverse G n) ·A (Group.inverse G m))) f (x ·A ((y ·A (Group.inverse G n)) ·A (Group.inverse G m)))
p3 : f (x ·A ((y ·A (Group.inverse G n)) ·A (Group.inverse G m))) (f x) ·B f ((y ·A (Group.inverse G n)) ·A (Group.inverse G m))
p4 : (f x) ·B f ((y ·A (Group.inverse G n)) ·A (Group.inverse G m)) (f x) ·B (f (y ·A (Group.inverse G n)) ·B f (Group.inverse G m))
p5 : (f x) ·B (f (y ·A (Group.inverse G n)) ·B f (Group.inverse G m)) (f x) ·B (Group.identity H ·B f (Group.inverse G m))
p6 : (f x) ·B (Group.identity H ·B f (Group.inverse G m)) (f x) ·B f (Group.inverse G m)
p7 : (f x) ·B f (Group.inverse G m) f (x ·A (Group.inverse G m))
p8 : f (x ·A (Group.inverse G m)) Group.identity H
p1 = GroupHom.wellDefined fHom (Group.wellDefined G (Reflexive.reflexive (Equivalence.reflexiveEq (Setoid.eq S))) (invContravariant G m n))
p2 = GroupHom.wellDefined fHom (Symmetric.symmetric (Equivalence.symmetricEq (Setoid.eq S)) (fourWayAssoc G))
p3 = GroupHom.groupHom fHom
p4 = Group.wellDefined H (Reflexive.reflexive reflexiveEq) (GroupHom.groupHom fHom)
p5 = Group.wellDefined H (Reflexive.reflexive reflexiveEq) (replaceGroupOp H (symmetric y~n) (Reflexive.reflexive reflexiveEq) (Reflexive.reflexive reflexiveEq) (Reflexive.reflexive reflexiveEq) (Reflexive.reflexive reflexiveEq))
p6 = Group.wellDefined H (Reflexive.reflexive reflexiveEq) (Group.multIdentLeft H)
p7 = Symmetric.symmetric symmetricEq (GroupHom.groupHom fHom)
p8 = x~m
ans : f ((x ·A y) ·A (Group.inverse G (m ·A n))) Group.identity H
ans = transitive p1 (transitive p2 (transitive p3 (transitive p4 (transitive p5 (transitive p6 (transitive p7 p8))))))
Group.identity (quotientGroup G fHom) = Group.identity G
Group.inverse (quotientGroup G fHom) = Group.inverse G
Group.multAssoc (quotientGroup {S = S} {T = T} {_·A_ = _·A_} G {H = H} {underf = f} fHom) {a} {b} {c} = ans
where
open Setoid T
open Equivalence (Setoid.eq T)
open Transitive transitiveEq
ans : f ((a ·A (b ·A c)) ·A (Group.inverse G ((a ·A b) ·A c))) Group.identity H
ans = transitive (GroupHom.wellDefined fHom (transferToRight'' G (Group.multAssoc G))) (imageOfIdentityIsIdentity fHom)
Group.multIdentRight (quotientGroup {S = S} {T = T} {_·A_ = _·A_} G {H = H} {underf = f} fHom) {a} = ans
where
open Group G
open Setoid T
open Equivalence eq
open Transitive transitiveEq
open Transitive (Equivalence.transitiveEq (Setoid.eq S)) renaming (transitive to transitiveG)
ans : f ((a ·A identity) ·A inverse a) Group.identity H
ans = transitive (GroupHom.wellDefined fHom (transitiveG (Group.wellDefined G (Group.multIdentRight G) (Reflexive.reflexive (Equivalence.reflexiveEq (Setoid.eq S)))) (Group.invRight G))) (imageOfIdentityIsIdentity fHom)
Group.multIdentLeft (quotientGroup {S = S} {T = T} {_·A_ = _·A_} G {H = H} {underf = f} fHom) {a} = ans
where
open Group G
open Setoid T
open Equivalence eq
open Transitive transitiveEq
open Transitive (Equivalence.transitiveEq (Setoid.eq S)) renaming (transitive to transitiveG)
ans : f ((identity ·A a) ·A (inverse a)) Group.identity H
ans = transitive (GroupHom.wellDefined fHom (transitiveG (Group.wellDefined G (Group.multIdentLeft G) (Reflexive.reflexive (Equivalence.reflexiveEq (Setoid.eq S)))) (Group.invRight G))) (imageOfIdentityIsIdentity fHom)
Group.invLeft (quotientGroup {S = S} {T = T} {_·A_ = _·A_} G {H = H} {underf = f} fHom) {x} = ans
where
open Group G
open Setoid T
open Equivalence eq
open Transitive transitiveEq
ans : f ((inverse x ·A x) ·A (inverse identity)) (Group.identity H)
ans = transitive (GroupHom.wellDefined fHom (Transitive.transitive (Equivalence.transitiveEq (Setoid.eq S)) (replaceGroupOp G (Symmetric.symmetric (Equivalence.symmetricEq (Setoid.eq S)) (Group.invLeft G)) (Symmetric.symmetric (Equivalence.symmetricEq (Setoid.eq S)) (invIdentity G)) (Reflexive.reflexive (Equivalence.reflexiveEq (Setoid.eq S))) ((Reflexive.reflexive (Equivalence.reflexiveEq (Setoid.eq S)))) ((Reflexive.reflexive (Equivalence.reflexiveEq (Setoid.eq S))))) (multIdentRight {identity}))) (imageOfIdentityIsIdentity fHom)
Group.invRight (quotientGroup {S = S} {T = T} {_·A_ = _·A_} G {H = H} {underf = f} fHom) {x} = ans
where
open Group G
open Setoid T
open Equivalence eq
open Transitive transitiveEq
ans : f ((x ·A inverse x) ·A (inverse identity)) (Group.identity H)
ans = transitive (GroupHom.wellDefined fHom (Transitive.transitive (Equivalence.transitiveEq (Setoid.eq S)) (replaceGroupOp G (Symmetric.symmetric (Equivalence.symmetricEq (Setoid.eq S)) (Group.invRight G)) (Symmetric.symmetric (Equivalence.symmetricEq (Setoid.eq S)) (invIdentity G)) (Reflexive.reflexive (Equivalence.reflexiveEq (Setoid.eq S))) (Reflexive.reflexive (Equivalence.reflexiveEq (Setoid.eq S))) (Reflexive.reflexive (Equivalence.reflexiveEq (Setoid.eq S)))) (multIdentRight {identity}))) (imageOfIdentityIsIdentity fHom)
record Subgroup {a} {b} {c} {d} {A : Set a} {B : Set c} {S : Setoid {a} {b} A} {T : Setoid {c} {d} B} {_·A_ : A A A} {_·B_ : B B B} (G : Group S _·A_) (H : Group T _·B_) {f : B A} (hom : GroupHom H G f) : Set (a b c d) where
open Setoid T renaming (__ to _G_)
open Setoid S renaming (__ to _H_)
field
fInj : SetoidInjection T S f
{-
quotientHom : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_·A_ : A → A → A} {_·B_ : B → B → B} (G : Group S _·A_) {H : Group T _·B_}{f : A → B} → (fHom : GroupHom G H f) → A → A
quotientHom {S = S} {T = T} {_·A_ = _·A_} {_·B_ = _·B_} G {f = f} fHom a = {!!}
quotientInjection : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_·A_ : A → A → A} {_·B_ : B → B → B} (G : Group S _·A_) {H : Group T _·B_}{f : A → B} → (fHom : GroupHom G H f) → GroupHom (quotientGroup G fHom) G (quotientHom G fHom)
GroupHom.groupHom (quotientInjection {S = S} {T = T} {_·A_ = _·A_} {_·B_ = _·B_} G {f = f} fHom) {x} {y} = {!!}
where
open Setoid S
open Equivalence eq
open Reflexive reflexiveEq
GroupHom.wellDefined (quotientInjection {S = S} {T = T} {_·A_ = _·A_} G {H = H} {f = f} fHom) {x} {y} x~y = {!!}
where
open Group G
open Setoid S
open Setoid T renaming (__ to _T_)
open Equivalence (Setoid.eq S)
open Reflexive reflexiveEq
have : f (x ·A inverse y) T Group.identity H
have = x~y
need : x y
need = {!!}
quotientIsSubgroup : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_·A_ : A → A → A} {_·B_ : B → B → B} {G : Group S _·A_} {H : Group T _·B_}{f : A → B}{fHom : GroupHom G H f} → Subgroup G (quotientGroup G fHom) (quotientInjection G fHom)
quotientIsSubgroup = {!!}
-}
subgroupOfAbelianIsAbelian : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+A_ : A A A} {_+B_ : B B B} {G : Group S _+A_} {H : Group T _+B_} {f : B A} {fHom : GroupHom H G f} Subgroup G H fHom AbelianGroup G AbelianGroup H
AbelianGroup.commutative (subgroupOfAbelianIsAbelian {S = S} {_+B_ = _+B_} {fHom = fHom} record { fInj = fInj } record { commutative = commutative }) {x} {y} = SetoidInjection.injective fInj (transitive (GroupHom.groupHom fHom) (transitive commutative (symmetric (GroupHom.groupHom fHom))))
where
open Setoid S
open Transitive (Equivalence.transitiveEq eq)
open Symmetric (Equivalence.symmetricEq eq)
abelianIsGroupProperty : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+A_ : A A A} {_+B_ : B B B} {G : Group S _+A_} {H : Group T _+B_} GroupsIsomorphic G H AbelianGroup H AbelianGroup G
abelianIsGroupProperty iso abH = subgroupOfAbelianIsAbelian {fHom = GroupIso.groupHom (GroupsIsomorphic.proof iso)} (record { fInj = SetoidBijection.inj (GroupIso.bij (GroupsIsomorphic.proof iso)) }) abH

174
Groups2.agda Normal file
View File

@@ -0,0 +1,174 @@
{-# OPTIONS --safe --warning=error #-}
open import Groups
open import Orders
open import Integers
open import Setoids
open import LogicalFormulae
open import FinSet
open import Functions
open import Naturals
open import IntegersModN
open import RingExamples
open import PrimeNumbers
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
module Groups2 where
data GroupHomImageElement {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+A_ : A A A} {_+B_ : B B B} {G : Group S _+A_} {H : Group T _+B_} {f : A B} (fHom : GroupHom G H f) : Set (a b c d) where
ofElt : (x : A) GroupHomImageElement fHom
imageGroupSetoid : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+A_ : A A A} {_+B_ : B B B} {G : Group S _+A_} {H : Group T _+B_} {f : A B} (fHom : GroupHom G H f) Setoid (GroupHomImageElement fHom)
(imageGroupSetoid {T = T} {f = f} fHom Setoid. ofElt b1) (ofElt b2) = Setoid.__ T (f b1) (f b2)
Reflexive.reflexive (Equivalence.reflexiveEq (Setoid.eq (imageGroupSetoid {T = T} fHom))) {ofElt b1} = Reflexive.reflexive (Equivalence.reflexiveEq (Setoid.eq T))
Symmetric.symmetric (Equivalence.symmetricEq (Setoid.eq (imageGroupSetoid {T = T} fHom))) {ofElt b1} {ofElt b2} = Symmetric.symmetric (Equivalence.symmetricEq (Setoid.eq T))
Transitive.transitive (Equivalence.transitiveEq (Setoid.eq (imageGroupSetoid {T = T} fHom))) {ofElt b1} {ofElt b2} {ofElt b3} = Transitive.transitive (Equivalence.transitiveEq (Setoid.eq T))
imageGroupOp : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+A_ : A A A} {_+B_ : B B B} {G : Group S _+A_} {H : Group T _+B_} {f : A B} (fHom : GroupHom G H f) GroupHomImageElement fHom GroupHomImageElement fHom GroupHomImageElement fHom
imageGroupOp {_+A_ = _+A_} fHom (ofElt a) (ofElt b) = ofElt (a +A b)
imageGroup : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+A_ : A A A} {_+B_ : B B B} {G : Group S _+A_} {H : Group T _+B_} {f : A B} (fHom : GroupHom G H f) Group (imageGroupSetoid fHom) (imageGroupOp fHom)
Group.wellDefined (imageGroup {T = T} {_+A_ = _+A_} {H = H} {f = f} fHom) {ofElt x} {ofElt y} {ofElt a} {ofElt b} x~a y~b = ans
where
open Setoid T
open Transitive (Equivalence.transitiveEq eq)
open Symmetric (Equivalence.symmetricEq eq)
ans : f (x +A y) f (a +A b)
ans = transitive (GroupHom.groupHom fHom) (transitive (Group.wellDefined H x~a y~b) (symmetric (GroupHom.groupHom fHom)))
Group.identity (imageGroup {G = G} fHom) = ofElt (Group.identity G)
Group.inverse (imageGroup {G = G} fHom) (ofElt a) = ofElt (Group.inverse G a)
Group.multAssoc (imageGroup {G = G} fHom) {ofElt a} {ofElt b} {ofElt c} = GroupHom.wellDefined fHom (Group.multAssoc G)
Group.multIdentRight (imageGroup {G = G} fHom) {ofElt a} = GroupHom.wellDefined fHom (Group.multIdentRight G)
Group.multIdentLeft (imageGroup {G = G} fHom) {ofElt a} = GroupHom.wellDefined fHom (Group.multIdentLeft G)
Group.invLeft (imageGroup {G = G} fHom) {ofElt a} = GroupHom.wellDefined fHom (Group.invLeft G)
Group.invRight (imageGroup {G = G} fHom) {ofElt a} = GroupHom.wellDefined fHom (Group.invRight G)
groupHomImageInclusion : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+G_ : A A A} {_+H_ : B B B} {G : Group S _+G_} {H : Group T _+H_} {f : A B} (fHom : GroupHom G H f) GroupHomImageElement fHom B
groupHomImageInclusion {f = f} fHom (ofElt x) = f x
groupHomImageIncludes : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+G_ : A A A} {_+H_ : B B B} {G : Group S _+G_} {H : Group T _+H_} {f : A B} (fHom : GroupHom G H f) GroupHom (imageGroup fHom) H (groupHomImageInclusion fHom)
GroupHom.groupHom (groupHomImageIncludes fHom) {ofElt x} {ofElt y} = GroupHom.groupHom fHom
GroupHom.wellDefined (groupHomImageIncludes fHom) {ofElt x} {ofElt y} x~y = x~y
groupHomImageIsSubgroup : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+G_ : A A A} {_+H_ : B B B} {G : Group S _+G_} {H : Group T _+H_} {f : A B} (fHom : GroupHom G H f) Subgroup H (imageGroup fHom) (groupHomImageIncludes fHom)
Subgroup.fInj (groupHomImageIsSubgroup {S = S} {T} {_+G_} {_+H_} {G} {H} {f} fHom) = record { wellDefined = λ {x} {y} GroupHom.wellDefined (groupHomImageIncludes fHom) {x} {y} ; injective = λ {x} {y} inj {x} {y} }
where
inj : {x y : GroupHomImageElement fHom} (Setoid.__ T (groupHomImageInclusion fHom x) (groupHomImageInclusion fHom y)) Setoid.__ (imageGroupSetoid fHom) x y
inj {ofElt x} {ofElt y} x~y = x~y
groupFirstIsomorphismIso : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+G_ : A A A} {_+H_ : B B B} {G : Group S _+G_} {H : Group T _+H_} {f : A B} (fHom : GroupHom G H f) GroupHomImageElement fHom A
groupFirstIsomorphismIso fHom (ofElt a) = a
groupFirstIsomorphismIsoHom : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+G_ : A A A} {_+H_ : B B B} {G : Group S _+G_} {H : Group T _+H_} {f : A B} (fHom : GroupHom G H f) GroupHom (imageGroup fHom) (quotientGroup G fHom) (groupFirstIsomorphismIso fHom)
GroupHom.groupHom (groupFirstIsomorphismIsoHom {G = G} fHom) {ofElt a} {ofElt b} = Reflexive.reflexive (Equivalence.reflexiveEq (Setoid.eq (quotientGroupSetoid G fHom)))
GroupHom.wellDefined (groupFirstIsomorphismIsoHom {T = T} {_+G_ = _+G_} {G = G} {H = H} {f = f} fHom) {ofElt a} {ofElt b} pr = ans
where
open Setoid T
open Transitive (Equivalence.transitiveEq (Setoid.eq T))
open Symmetric (Equivalence.symmetricEq (Setoid.eq T))
open Reflexive (Equivalence.reflexiveEq (Setoid.eq T))
ans : f (a +G Group.inverse G b) Group.identity H
ans = transitive (GroupHom.groupHom fHom) (transitive (Group.wellDefined H pr reflexive) (transitive (symmetric (GroupHom.groupHom fHom)) (transitive (GroupHom.wellDefined fHom (Group.invRight G)) (imageOfIdentityIsIdentity fHom))))
groupFirstIsomorphismTheorem' : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+G_ : A A A} {_+H_ : B B B} {G : Group S _+G_} {H : Group T _+H_} {f : A B} (fHom : GroupHom G H f) GroupIso (imageGroup fHom) (quotientGroup G fHom) (groupFirstIsomorphismIso fHom)
GroupIso.groupHom (groupFirstIsomorphismTheorem' fHom) = groupFirstIsomorphismIsoHom fHom
SetoidInjection.wellDefined (SetoidBijection.inj (GroupIso.bij (groupFirstIsomorphismTheorem' fHom))) {x} {y} x~y = GroupHom.wellDefined (groupFirstIsomorphismIsoHom fHom) {x} {y} x~y
SetoidInjection.injective (SetoidBijection.inj (GroupIso.bij (groupFirstIsomorphismTheorem' {T = T} {H = H} {f = f} fHom))) {ofElt a} {ofElt b} pr = need
where
open Setoid T
open Transitive (Equivalence.transitiveEq (Setoid.eq T))
open Symmetric (Equivalence.symmetricEq (Setoid.eq T))
open Reflexive (Equivalence.reflexiveEq (Setoid.eq T))
need : f a f b
need = transferToRight H (transitive (transitive (Group.wellDefined H reflexive (symmetric (homRespectsInverse fHom))) (symmetric (GroupHom.groupHom fHom))) pr)
SetoidSurjection.wellDefined (SetoidBijection.surj (GroupIso.bij (groupFirstIsomorphismTheorem' fHom))) {x} {y} x~y = GroupHom.wellDefined (groupFirstIsomorphismIsoHom fHom) {x} {y} x~y
SetoidSurjection.surjective (SetoidBijection.surj (GroupIso.bij (groupFirstIsomorphismTheorem' {G = G} fHom))) {a} = ofElt a , Reflexive.reflexive (Equivalence.reflexiveEq (Setoid.eq (quotientGroupSetoid G fHom)))
groupFirstIsomorphismTheorem : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+G_ : A A A} {_+H_ : B B B} {G : Group S _+G_} {H : Group T _+H_} {f : A B} (fHom : GroupHom G H f) GroupsIsomorphic (imageGroup fHom) (quotientGroup G fHom)
groupFirstIsomorphismTheorem fHom = record { isomorphism = groupFirstIsomorphismIso fHom ; proof = groupFirstIsomorphismTheorem' fHom }
record NormalSubgroup {a} {b} {c} {d} {A : Set a} {B : Set c} {S : Setoid {a} {b} A} {T : Setoid {c} {d} B} {_·A_ : A A A} {_·B_ : B B B} (G : Group S _·A_) (H : Group T _·B_) {f : B A} (hom : GroupHom H G f) : Set (a b c d) where
open Setoid S
field
subgroup : Subgroup G H hom
normal : {g : A} {h : B} Sg B (λ fromH (g ·A (f h)) ·A (Group.inverse G g) f fromH)
data GroupKernelElement {a} {b} {c} {d} {A : Set a} {B : Set c} {S : Setoid {a} {b} A} {T : Setoid {c} {d} B} {_·A_ : A A A} {_·B_ : B B B} (G : Group S _·A_) {H : Group T _·B_} {f : A B} (hom : GroupHom G H f) : Set (a b c d) where
kerOfElt : (x : A) (Setoid.__ T (f x) (Group.identity H)) GroupKernelElement G hom
groupKernel : {a b c d : _} {A : Set a} {B : Set c} {S : Setoid {a} {b} A} {T : Setoid {c} {d} B} {_·A_ : A A A} {_·B_ : B B B} (G : Group S _·A_) {H : Group T _·B_} {f : A B} (hom : GroupHom G H f) Setoid (GroupKernelElement G hom)
Setoid.__ (groupKernel {S = S} G {H} {f} fHom) (kerOfElt x fx=0) (kerOfElt y fy=0) = Setoid.__ S x y
Reflexive.reflexive (Equivalence.reflexiveEq (Setoid.eq (groupKernel {S = S} G {H} {f} fHom))) {kerOfElt x x₁} = Reflexive.reflexive (Equivalence.reflexiveEq (Setoid.eq S))
Symmetric.symmetric (Equivalence.symmetricEq (Setoid.eq (groupKernel {S = S} G {H} {f} fHom))) {kerOfElt x prX} {kerOfElt y prY} = Symmetric.symmetric (Equivalence.symmetricEq (Setoid.eq S))
Transitive.transitive (Equivalence.transitiveEq (Setoid.eq (groupKernel {S = S} G {H} {f} fHom))) {kerOfElt x prX} {kerOfElt y prY} {kerOfElt z prZ} = Transitive.transitive (Equivalence.transitiveEq (Setoid.eq S))
groupKernelGroupOp : {a b c d : _} {A : Set a} {B : Set c} {S : Setoid {a} {b} A} {T : Setoid {c} {d} B} {_·A_ : A A A} {_·B_ : B B B} (G : Group S _·A_) {H : Group T _·B_} {f : A B} (hom : GroupHom G H f) (GroupKernelElement G hom) (GroupKernelElement G hom) (GroupKernelElement G hom)
groupKernelGroupOp {T = T} {_·A_ = _+A_} G {H = H} hom (kerOfElt x prX) (kerOfElt y prY) = kerOfElt (x +A y) (transitive (GroupHom.groupHom hom) (transitive (Group.wellDefined H prX prY) (Group.multIdentLeft H)))
where
open Setoid T
open Transitive (Equivalence.transitiveEq (Setoid.eq T))
groupKernelGroup : {a b c d : _} {A : Set a} {B : Set c} {S : Setoid {a} {b} A} {T : Setoid {c} {d} B} {_·A_ : A A A} {_·B_ : B B B} (G : Group S _·A_) {H : Group T _·B_} {f : A B} (hom : GroupHom G H f) Group (groupKernel G hom) (groupKernelGroupOp G hom)
Group.wellDefined (groupKernelGroup G fHom) {kerOfElt x prX} {kerOfElt y prY} {kerOfElt a prA} {kerOfElt b prB} = Group.wellDefined G
Group.identity (groupKernelGroup G fHom) = kerOfElt (Group.identity G) (imageOfIdentityIsIdentity fHom)
Group.inverse (groupKernelGroup {T = T} G {H = H} fHom) (kerOfElt x prX) = kerOfElt (Group.inverse G x) (transitive (homRespectsInverse fHom) (transitive (inverseWellDefined H prX) (invIdentity H)))
where
open Setoid T
open Transitive (Equivalence.transitiveEq (Setoid.eq T))
Group.multAssoc (groupKernelGroup {S = S} {_·A_ = _·A_} G fHom) {kerOfElt x prX} {kerOfElt y prY} {kerOfElt z prZ} = Group.multAssoc G
Group.multIdentRight (groupKernelGroup G fHom) {kerOfElt x prX} = Group.multIdentRight G
Group.multIdentLeft (groupKernelGroup G fHom) {kerOfElt x prX} = Group.multIdentLeft G
Group.invLeft (groupKernelGroup G fHom) {kerOfElt x prX} = Group.invLeft G
Group.invRight (groupKernelGroup G fHom) {kerOfElt x prX} = Group.invRight G
injectionFromKernelToG : {a b c d : _} {A : Set a} {B : Set c} {S : Setoid {a} {b} A} {T : Setoid {c} {d} B} {_·A_ : A A A} {_·B_ : B B B} (G : Group S _·A_) {H : Group T _·B_} {f : A B} (hom : GroupHom G H f) GroupKernelElement G hom A
injectionFromKernelToG G hom (kerOfElt x _) = x
injectionFromKernelToGIsHom : {a b c d : _} {A : Set a} {B : Set c} {S : Setoid {a} {b} A} {T : Setoid {c} {d} B} {_·A_ : A A A} {_·B_ : B B B} (G : Group S _·A_) {H : Group T _·B_} {f : A B} (hom : GroupHom G H f) GroupHom (groupKernelGroup G hom) G (injectionFromKernelToG G hom)
GroupHom.groupHom (injectionFromKernelToGIsHom {S = S} G hom) {kerOfElt x prX} {kerOfElt y prY} = Reflexive.reflexive (Equivalence.reflexiveEq (Setoid.eq S))
GroupHom.wellDefined (injectionFromKernelToGIsHom G hom) {kerOfElt x prX} {kerOfElt y prY} i = i
groupKernelGroupIsSubgroup : {a b c d : _} {A : Set a} {B : Set c} {S : Setoid {a} {b} A} {T : Setoid {c} {d} B} {_·A_ : A A A} {_·B_ : B B B} (G : Group S _·A_) {H : Group T _·B_} {f : A B} (hom : GroupHom G H f) Subgroup G (groupKernelGroup G hom) (injectionFromKernelToGIsHom G hom)
Subgroup.fInj (groupKernelGroupIsSubgroup {S = S} {T = T} G {f = f} hom) = record { wellDefined = λ {x} {y} GroupHom.wellDefined (injectionFromKernelToGIsHom G hom) {x} {y} ; injective = λ {x} {y} inj {x} {y} }
where
inj : {x : GroupKernelElement G hom} {y : GroupKernelElement G hom} Setoid.__ S (injectionFromKernelToG G hom x) (injectionFromKernelToG G hom y) Setoid.__ (groupKernel G hom) x y
inj {kerOfElt x prX} {kerOfElt y prY} = id
groupKernelGroupIsNormalSubgroup : {a b c d : _} {A : Set a} {B : Set c} {S : Setoid {a} {b} A} {T : Setoid {c} {d} B} {_·A_ : A A A} {_·B_ : B B B} (G : Group S _·A_) {H : Group T _·B_} {f : A B} (hom : GroupHom G H f) NormalSubgroup G (groupKernelGroup G hom) (injectionFromKernelToGIsHom G hom)
NormalSubgroup.subgroup (groupKernelGroupIsNormalSubgroup G hom) = groupKernelGroupIsSubgroup G hom
NormalSubgroup.normal (groupKernelGroupIsNormalSubgroup {S = S} {T = T} {_·A_ = _·A_} G {H = H} {f = f} hom) {g} {kerOfElt h prH} = kerOfElt ((g ·A h) ·A Group.inverse G g) ans , Reflexive.reflexive (Equivalence.reflexiveEq (Setoid.eq S))
where
open Setoid T
open Transitive (Equivalence.transitiveEq (Setoid.eq T))
open Reflexive (Equivalence.reflexiveEq (Setoid.eq T))
open Symmetric (Equivalence.symmetricEq (Setoid.eq T))
ans : f ((g ·A h) ·A Group.inverse G g) Group.identity H
ans = transitive (GroupHom.groupHom hom) (transitive (Group.wellDefined H (GroupHom.groupHom hom) reflexive) (transitive (Group.wellDefined H (Group.wellDefined H reflexive prH) reflexive) (transitive (Group.wellDefined H (Group.multIdentRight H) reflexive) (transitive (symmetric (GroupHom.groupHom hom)) (transitive (GroupHom.wellDefined hom (Group.invRight G)) (imageOfIdentityIsIdentity hom))))))
abelianGroupSubgroupIsNormal : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+A_ : A A A} {_+B_ : B B B} {underG : Group S _+A_} (G : AbelianGroup underG) {H : Group T _+B_} {f : B A} {hom : GroupHom H underG f} (s : Subgroup underG H hom) NormalSubgroup underG H hom
NormalSubgroup.subgroup (abelianGroupSubgroupIsNormal G H) = H
NormalSubgroup.normal (abelianGroupSubgroupIsNormal {S = S} {underG = G} record { commutative = commutative } H) {g} {h} = h , transitive (wellDefined commutative reflexive) (transitive (symmetric multAssoc) (transitive (wellDefined reflexive invRight) multIdentRight))
where
open Setoid S
open Group G
open Reflexive (Equivalence.reflexiveEq (Setoid.eq S))
open Transitive (Equivalence.transitiveEq (Setoid.eq S))
open Symmetric (Equivalence.symmetricEq (Setoid.eq S))
trivialGroup : Group (reflSetoid (FinSet 1)) λ _ _ fzero
Group.wellDefined trivialGroup _ _ = refl
Group.identity trivialGroup = fzero
Group.inverse trivialGroup _ = fzero
Group.multAssoc trivialGroup = refl
Group.multIdentRight trivialGroup {fzero} = refl
Group.multIdentRight trivialGroup {fsucc ()}
Group.multIdentLeft trivialGroup {fzero} = refl
Group.multIdentLeft trivialGroup {fsucc ()}
Group.invLeft trivialGroup = refl
Group.invRight trivialGroup = refl
identityHom : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+A_ : A A A} (G : Group S _+A_) GroupHom G G id
GroupHom.groupHom (identityHom {S = S} G) = Reflexive.reflexive (Equivalence.reflexiveEq (Setoid.eq S))
GroupHom.wellDefined (identityHom G) = id

250
GroupsExampleSheet1.agda Normal file
View File

@@ -0,0 +1,250 @@
{-# OPTIONS --safe --warning=error #-}
open import LogicalFormulae
open import Setoids
open import Functions
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Naturals
open import Integers
open import FinSet
open import Groups
open import Rings
open import Fields
module GroupsExampleSheet1 where
{-
Question 1: e is the unique solution of x^2 = x
-}
question1 : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A A A} (G : Group S _+_) (x : A) Setoid.__ S (x + x) x Setoid.__ S x (Group.identity G)
question1 {S = S} {_+_ = _+_} G x x+x=x = transitive (symmetric multIdentRight) (transitive (wellDefined reflexive (symmetric invRight)) (transitive multAssoc (transitive (wellDefined x+x=x reflexive) invRight)))
where
open Group G
open Setoid S
open Transitive (Equivalence.transitiveEq (Setoid.eq S))
open Reflexive (Equivalence.reflexiveEq (Setoid.eq S))
open Symmetric (Equivalence.symmetricEq (Setoid.eq S))
question1' : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A A A} (G : Group S _+_) Setoid.__ S ((Group.identity G) + (Group.identity G)) (Group.identity G)
question1' G = Group.multIdentRight G
{-
Question 2: intersection of subgroups is a subgroup; union of subgroups is a subgroup iff one is contained in the other.
First, define the intersection of subgroups and show that it is a subgroup.
-}
data SubgroupIntersectionElement {a b c d e f : _} {A : Set a} {B : Set b} {C : Set c} {S : Setoid {a} {d} A} {T : Setoid {b} {e} B} {U : Setoid {c} {f} C} {_+_ : A A A} {_+H1_ : B B B} {_+H2_ : C C C} (G : Group S _+_) {H1grp : Group T _+H1_} {H2grp : Group U _+H2_} {h1Inj : B A} {h2Inj : C A} {h1Hom : GroupHom H1grp G h1Inj} {h2Hom : GroupHom H2grp G h2Inj} (H1 : Subgroup G H1grp h1Hom) (H2 : Subgroup G H2grp h2Hom) : Set (a b c d e f) where
ofElt : {x : A} Sg B (λ b Setoid.__ S (h1Inj b) x) Sg C (λ c Setoid.__ S (h2Inj c) x) SubgroupIntersectionElement G H1 H2
subgroupIntersectionOp : {a b c d e f : _} {A : Set a} {B : Set b} {C : Set c} {S : Setoid {a} {d} A} {T : Setoid {b} {e} B} {U : Setoid {c} {f} C} {_+_ : A A A} {_+H1_ : B B B} {_+H2_ : C C C} (G : Group S _+_) {H1grp : Group T _+H1_} {H2grp : Group U _+H2_} {h1Inj : B A} {h2Inj : C A} {h1Hom : GroupHom H1grp G h1Inj} {h2Hom : GroupHom H2grp G h2Inj} (H1 : Subgroup G H1grp h1Hom) (H2 : Subgroup G H2grp h2Hom) (r : SubgroupIntersectionElement G H1 H2) (s : SubgroupIntersectionElement G H1 H2) SubgroupIntersectionElement G H1 H2
subgroupIntersectionOp {S = S} {_+_ = _+_} {_+H1_ = _+H1_} {_+H2_ = _+H2_} G {h1Hom = h1Hom} {h2Hom = h2Hom} H1 H2 (ofElt (b , prB) (c , prC)) (ofElt (b2 , prB2) (c2 , prC2)) = ofElt ((b +H1 b2) , GroupHom.groupHom h1Hom) ((c +H2 c2) , transitive (GroupHom.groupHom h2Hom) (transitive (Group.wellDefined G prC prC2) (Group.wellDefined G (symmetric prB) (symmetric prB2))))
where
open Setoid S
open Transitive (Equivalence.transitiveEq (Setoid.eq S))
open Symmetric (Equivalence.symmetricEq (Setoid.eq S))
subgroupIntersectionSetoid : {a b c d e f : _} {A : Set a} {B : Set b} {C : Set c} {S : Setoid {a} {d} A} {T : Setoid {b} {e} B} {U : Setoid {c} {f} C} {_+_ : A A A} {_+H1_ : B B B} {_+H2_ : C C C} (G : Group S _+_) {H1grp : Group T _+H1_} {H2grp : Group U _+H2_} {h1Inj : B A} {h2Inj : C A} {h1Hom : GroupHom H1grp G h1Inj} {h2Hom : GroupHom H2grp G h2Inj} (H1 : Subgroup G H1grp h1Hom) (H2 : Subgroup G H2grp h2Hom) Setoid (SubgroupIntersectionElement G H1 H2)
Setoid.__ (subgroupIntersectionSetoid {T = T} {U = U} G {h1Inj = h1} {h2Inj = h2} H1 H2) (ofElt (xH1 , prxH1) (xH2 , prxH2)) (ofElt (yH1 , pryH1) (yH2 , pryH2)) = (Setoid.__ T xH1 yH1) && (Setoid.__ U xH2 yH2)
Reflexive.reflexive (Equivalence.reflexiveEq (Setoid.eq (subgroupIntersectionSetoid {T = T} {U = U} G H1 H2))) {ofElt (a , prA) (b , prB)} = (Reflexive.reflexive (Equivalence.reflexiveEq (Setoid.eq T))) ,, (Reflexive.reflexive (Equivalence.reflexiveEq (Setoid.eq U)))
Symmetric.symmetric (Equivalence.symmetricEq (Setoid.eq (subgroupIntersectionSetoid {T = T} {U = U} G H1 H2))) {ofElt (a , prA) (b , prB)} {ofElt (c , prC) (d , prD)} (fst ,, snd) = Symmetric.symmetric (Equivalence.symmetricEq (Setoid.eq T)) fst ,, Symmetric.symmetric (Equivalence.symmetricEq (Setoid.eq U)) snd
Transitive.transitive (Equivalence.transitiveEq (Setoid.eq (subgroupIntersectionSetoid {T = T} {U = U} G H1 H2))) {ofElt (_ , _) (_ , _)} {ofElt (_ , _) (_ , _)} {ofElt (_ , _) (_ , _)} (fst1 ,, snd1) (fst2 ,, snd2) = Transitive.transitive (Equivalence.transitiveEq (Setoid.eq T)) fst1 fst2 ,, Transitive.transitive (Equivalence.transitiveEq (Setoid.eq U)) snd1 snd2
subgroupIntersectionGroup : {a b c d e f : _} {A : Set a} {B : Set b} {C : Set c} {S : Setoid {a} {d} A} {T : Setoid {b} {e} B} {U : Setoid {c} {f} C} {_+_ : A A A} {_+H1_ : B B B} {_+H2_ : C C C} (G : Group S _+_) {H1grp : Group T _+H1_} {H2grp : Group U _+H2_} {h1Inj : B A} {h2Inj : C A} {h1Hom : GroupHom H1grp G h1Inj} {h2Hom : GroupHom H2grp G h2Inj} (H1 : Subgroup G H1grp h1Hom) (H2 : Subgroup G H2grp h2Hom) Group (subgroupIntersectionSetoid G H1 H2) (subgroupIntersectionOp G H1 H2)
Group.wellDefined (subgroupIntersectionGroup {S = S} {T = T} {U = U} G {H1grp = h1} {H2grp = h2} H1 H2) {ofElt (_ , _) (_ , _)} {ofElt (_ , _ ) (_ , _)} {ofElt (_ , _) (_ , _)} {ofElt (_ , _) (_ , _)} (pr1 ,, pr2) (pr3 ,, pr4) = transitiveT (Group.wellDefined h1 pr1 reflexiveT) (Group.wellDefined h1 reflexiveT pr3) ,, transitiveU (Group.wellDefined h2 pr2 reflexiveU) ((Group.wellDefined h2 reflexiveU pr4))
where
open Group G
open Setoid T
open Transitive (Equivalence.transitiveEq (Setoid.eq T)) renaming (transitive to transitiveT)
open Symmetric (Equivalence.symmetricEq (Setoid.eq T)) renaming (symmetric to symmetricT)
open Reflexive (Equivalence.reflexiveEq (Setoid.eq T)) renaming (reflexive to reflexiveT)
open Transitive (Equivalence.transitiveEq (Setoid.eq U)) renaming (transitive to transitiveU)
open Symmetric (Equivalence.symmetricEq (Setoid.eq U)) renaming (symmetric to symmetricU)
open Reflexive (Equivalence.reflexiveEq (Setoid.eq U)) renaming (reflexive to reflexiveU)
Group.identity (subgroupIntersectionGroup G {H1grp = H1grp} {H2grp = H2grp} {h1Hom = h1Hom} {h2Hom = h2Hom} H1 H2) = ofElt {x = Group.identity G} (Group.identity H1grp , imageOfIdentityIsIdentity h1Hom) (Group.identity H2grp , imageOfIdentityIsIdentity h2Hom)
Group.inverse (subgroupIntersectionGroup {S = S} G {H1grp = h1} {H2grp = h2} {h1Hom = h1hom} {h2Hom = h2hom} H1 H2) (ofElt (a , prA) (b , prB)) = ofElt (Group.inverse h1 a , homRespectsInverse h1hom) (Group.inverse h2 b , transitive (homRespectsInverse h2hom) (inverseWellDefined G (transitive prB (symmetric prA))))
where
open Setoid S
open Transitive (Equivalence.transitiveEq (Setoid.eq S))
open Symmetric (Equivalence.symmetricEq (Setoid.eq S))
Group.multAssoc (subgroupIntersectionGroup G {H1grp = h1} {H2grp = h2} H1 H2) {ofElt (a , prA) (b , prB)} {ofElt (c , prC) (d , prD)} {ofElt (e , prE) (f , prF)} = Group.multAssoc h1 ,, Group.multAssoc h2
Group.multIdentRight (subgroupIntersectionGroup G {H1grp = h1} {H2grp = h2} H1 H2) {ofElt (_ , _) (_ , _)} = Group.multIdentRight h1 ,, Group.multIdentRight h2
Group.multIdentLeft (subgroupIntersectionGroup G {H1grp = h1} {H2grp = h2} H1 H2) {ofElt (_ , _) (_ , _)} = Group.multIdentLeft h1 ,, Group.multIdentLeft h2
Group.invLeft (subgroupIntersectionGroup G {H1grp = h1} {H2grp = h2} H1 H2) {ofElt (_ , _) (_ , _)} = Group.invLeft h1 ,, Group.invLeft h2
Group.invRight (subgroupIntersectionGroup G {H1grp = h1} {H2grp = h2} H1 H2) {ofElt (_ , _) (_ , _)} = Group.invRight h1 ,, Group.invRight h2
subgroupIntersectionInjectionIntoMain : {a b c d e f : _} {A : Set a} {B : Set b} {C : Set c} {S : Setoid {a} {d} A} {T : Setoid {b} {e} B} {U : Setoid {c} {f} C} {_+_ : A A A} {_+H1_ : B B B} {_+H2_ : C C C} (G : Group S _+_) {H1grp : Group T _+H1_} {H2grp : Group U _+H2_} {h1Inj : B A} {h2Inj : C A} {h1Hom : GroupHom H1grp G h1Inj} {h2Hom : GroupHom H2grp G h2Inj} (H1 : Subgroup G H1grp h1Hom) (H2 : Subgroup G H2grp h2Hom) SubgroupIntersectionElement G H1 H2 A
subgroupIntersectionInjectionIntoMain G {h1Inj = f} H1 H2 (ofElt (a , prA) (b , prB)) = f a
subgroupIntersectionInjectionIntoMainIsHom : {a b c d e f : _} {A : Set a} {B : Set b} {C : Set c} {S : Setoid {a} {d} A} {T : Setoid {b} {e} B} {U : Setoid {c} {f} C} {_+_ : A A A} {_+H1_ : B B B} {_+H2_ : C C C} (G : Group S _+_) {H1grp : Group T _+H1_} {H2grp : Group U _+H2_} {h1Inj : B A} {h2Inj : C A} {h1Hom : GroupHom H1grp G h1Inj} {h2Hom : GroupHom H2grp G h2Inj} (H1 : Subgroup G H1grp h1Hom) (H2 : Subgroup G H2grp h2Hom) GroupHom (subgroupIntersectionGroup G H1 H2) G (subgroupIntersectionInjectionIntoMain G H1 H2)
GroupHom.groupHom (subgroupIntersectionInjectionIntoMainIsHom G {h1Hom = h1} H1 H2) {ofElt (_ , _) (_ , _)} {ofElt (_ , _) (_ , _)} = GroupHom.groupHom h1
GroupHom.wellDefined (subgroupIntersectionInjectionIntoMainIsHom G {h1Hom = h1} H1 H2) {ofElt (_ , _) (_ , _)} {ofElt (_ , _) (_ , _)} (fst ,, snd) = GroupHom.wellDefined h1 fst
subgroupIntersectionIsSubgroup : {a b c d e f : _} {A : Set a} {B : Set b} {C : Set c} {S : Setoid {a} {d} A} {T : Setoid {b} {e} B} {U : Setoid {c} {f} C} {_+_ : A A A} {_+H1_ : B B B} {_+H2_ : C C C} (G : Group S _+_) {H1grp : Group T _+H1_} {H2grp : Group U _+H2_} {h1Inj : B A} {h2Inj : C A} {h1Hom : GroupHom H1grp G h1Inj} {h2Hom : GroupHom H2grp G h2Inj} (H1 : Subgroup G H1grp h1Hom) (H2 : Subgroup G H2grp h2Hom) Subgroup G (subgroupIntersectionGroup G H1 H2) (subgroupIntersectionInjectionIntoMainIsHom G H1 H2)
SetoidInjection.wellDefined (Subgroup.fInj (subgroupIntersectionIsSubgroup G {h1Hom = h1} H1 H2)) {ofElt (_ , _) (_ , _)} {ofElt (_ , _) (_ , _)} (fst ,, snd) = GroupHom.wellDefined h1 fst
SetoidInjection.injective (Subgroup.fInj (subgroupIntersectionIsSubgroup {S = S} G H1 H2)) {ofElt (a , prA) (b , prB)} {ofElt (c , prC) (d , prD)} x~y = SetoidInjection.injective (Subgroup.fInj H1) x~y ,, SetoidInjection.injective (Subgroup.fInj H2) (transitive prB (transitive (transitive (symmetric prA) (transitive x~y prC) ) (symmetric prD)))
where
open Setoid S
open Transitive (Equivalence.transitiveEq (Setoid.eq S))
open Symmetric (Equivalence.symmetricEq (Setoid.eq S))
{-
To make sure we haven't defined something stupid, check that the intersection doesn't care which order the two subgroups came in, and check that the subgroup intersection is isomorphic to the original group in the case that the two were the same.
-}
subgroupIntersectionIsomorphic : {a b c d e f : _} {A : Set a} {B : Set b} {C : Set c} {S : Setoid {a} {d} A} {T : Setoid {b} {e} B} {U : Setoid {c} {f} C} {_+_ : A A A} {_+H1_ : B B B} {_+H2_ : C C C} (G : Group S _+_) {H1grp : Group T _+H1_} {H2grp : Group U _+H2_} {h1Inj : B A} {h2Inj : C A} {h1Hom : GroupHom H1grp G h1Inj} {h2Hom : GroupHom H2grp G h2Inj} (H1 : Subgroup G H1grp h1Hom) (H2 : Subgroup G H2grp h2Hom) GroupsIsomorphic (subgroupIntersectionGroup G H1 H2) (subgroupIntersectionGroup G H2 H1)
GroupsIsomorphic.isomorphism (subgroupIntersectionIsomorphic G H1 H2) (ofElt (a , prA) (b , prB)) = ofElt (b , prB) (a , prA)
GroupHom.groupHom (GroupIso.groupHom (GroupsIsomorphic.proof (subgroupIntersectionIsomorphic {T = T} {U = U} G H1 H2))) {ofElt (_ , _) (_ , _)} {ofElt (_ , _) (_ , _)} = Reflexive.reflexive (Equivalence.reflexiveEq (Setoid.eq U)) ,, Reflexive.reflexive (Equivalence.reflexiveEq (Setoid.eq T))
GroupHom.wellDefined (GroupIso.groupHom (GroupsIsomorphic.proof (subgroupIntersectionIsomorphic G H1 H2))) {ofElt (a , prA) (b , prB)} {ofElt (_ , _) (_ , _)} (fst ,, snd) = snd ,, fst
SetoidInjection.wellDefined (SetoidBijection.inj (GroupIso.bij (GroupsIsomorphic.proof (subgroupIntersectionIsomorphic G H1 H2)))) {ofElt (_ , _) (_ , _)} {ofElt (_ , _) (_ , _)} (fst ,, snd) = snd ,, fst
SetoidInjection.injective (SetoidBijection.inj (GroupIso.bij (GroupsIsomorphic.proof (subgroupIntersectionIsomorphic G H1 H2)))) {ofElt (_ , _) (_ , _)} {ofElt (_ , _) (_ , _)} (fst ,, snd) = snd ,, fst
SetoidSurjection.wellDefined (SetoidBijection.surj (GroupIso.bij (GroupsIsomorphic.proof (subgroupIntersectionIsomorphic G H1 H2)))) {ofElt (_ , _) (_ , _)} {ofElt (_ , _) (_ , _)} (fst ,, snd) = snd ,, fst
SetoidSurjection.surjective (SetoidBijection.surj (GroupIso.bij (GroupsIsomorphic.proof (subgroupIntersectionIsomorphic {T = T} {U = U} G H1 H2)))) {ofElt (a , prA) (b , prB)} = ofElt (b , prB) (a , prA) , (Reflexive.reflexive (Equivalence.reflexiveEq (Setoid.eq U)) ,, Reflexive.reflexive (Equivalence.reflexiveEq (Setoid.eq T)))
subgroupIntersectionOfSame : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+_ : A A A} {_+H1_ : B B B} (G : Group S _+_) {H1grp : Group T _+H1_} {h1Inj : B A} {h1Hom : GroupHom H1grp G h1Inj} (H1 : Subgroup G H1grp h1Hom) GroupsIsomorphic (subgroupIntersectionGroup G H1 H1) H1grp
GroupsIsomorphic.isomorphism (subgroupIntersectionOfSame G H1) (ofElt (a , prA) (b , prB)) = a
GroupHom.groupHom (GroupIso.groupHom (GroupsIsomorphic.proof (subgroupIntersectionOfSame {T = T} G H1))) {ofElt (_ , _) (_ , _)} {ofElt (_ , _) (_ , _)} = Reflexive.reflexive (Equivalence.reflexiveEq (Setoid.eq T))
GroupHom.wellDefined (GroupIso.groupHom (GroupsIsomorphic.proof (subgroupIntersectionOfSame G H1))) {ofElt (_ , _) (_ , _)} {ofElt (_ , _) (_ , _)} (fst ,, _) = fst
SetoidInjection.wellDefined (SetoidBijection.inj (GroupIso.bij (GroupsIsomorphic.proof (subgroupIntersectionOfSame G H1)))) {ofElt (_ , _) (_ , _)} {ofElt (_ , _) (_ , _)} (fst ,, _) = fst
SetoidInjection.injective (SetoidBijection.inj (GroupIso.bij (GroupsIsomorphic.proof (subgroupIntersectionOfSame {S = S} {T = T} G {h1Hom = h1Hom} H1)))) {ofElt (a , prA) (b , prB)} {ofElt (c , prC) (d , prD)} a~b = a~b ,, SetoidInjection.injective (Subgroup.fInj H1) (transitive prB (transitive (transitive (symmetric prA) (transitive (GroupHom.wellDefined h1Hom a~b) prC)) (symmetric prD)))
where
open Setoid S
open Transitive (Equivalence.transitiveEq (Setoid.eq S))
open Symmetric (Equivalence.symmetricEq (Setoid.eq S))
SetoidSurjection.wellDefined (SetoidBijection.surj (GroupIso.bij (GroupsIsomorphic.proof (subgroupIntersectionOfSame G H1)))) {ofElt (_ , _) (_ , _)} {ofElt (_ , _) (_ , _)} (fst ,, _) = fst
SetoidSurjection.surjective (SetoidBijection.surj (GroupIso.bij (GroupsIsomorphic.proof (subgroupIntersectionOfSame {S = S} {T = T} G H1)))) {b} = ofElt (b , Reflexive.reflexive (Equivalence.reflexiveEq (Setoid.eq S))) (b , Reflexive.reflexive (Equivalence.reflexiveEq (Setoid.eq S))) , (Reflexive.reflexive (Equivalence.reflexiveEq (Setoid.eq T)))
{- TODO: finish question 2 -}
{-
Question 3. We can't talk about yet, so we'll just work in an arbitrary integral domain.
Show that the collection of linear functions over a ring forms a group; is it abelian?
-}
record LinearFunction {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A A A} {_*_ : A A A} {R : Ring S _+_ _*_} (F : Field R) : Set (a b) where
field
xCoeff : A
xCoeffNonzero : (Setoid.__ S xCoeff (Ring.0R R) False)
constant : A
interpretLinearFunction : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A A A} {_*_ : A A A} {R : Ring S _+_ _*_} {F : Field R} (f : LinearFunction F) A A
interpretLinearFunction {_+_ = _+_} {_*_ = _*_} record { xCoeff = xCoeff ; xCoeffNonzero = xCoeffNonzero ; constant = constant } a = (xCoeff * a) + constant
composeLinearFunctions : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A A A} {_*_ : A A A} {R : Ring S _+_ _*_} {F : Field R} (f1 : LinearFunction F) (f2 : LinearFunction F) LinearFunction F
LinearFunction.xCoeff (composeLinearFunctions {_+_ = _+_} {_*_ = _*_} record { xCoeff = xCoeff1 ; xCoeffNonzero = xCoeffNonzero1 ; constant = constant1 } record { xCoeff = xCoeff2 ; xCoeffNonzero = xCoeffNonzero2 ; constant = constant2 }) = xCoeff1 * xCoeff2
LinearFunction.xCoeffNonzero (composeLinearFunctions {S = S} {R = R} {F = F} record { xCoeff = xCoeff1 ; xCoeffNonzero = xCoeffNonzero1 ; constant = constant1 } record { xCoeff = xCoeff2 ; xCoeffNonzero = xCoeffNonzero2 ; constant = constant2 }) pr = xCoeffNonzero2 bad
where
open Setoid S
open Ring R
open Transitive (Equivalence.transitiveEq (Setoid.eq S))
open Reflexive (Equivalence.reflexiveEq (Setoid.eq S))
open Symmetric (Equivalence.symmetricEq (Setoid.eq S))
bad : Setoid.__ S xCoeff2 0R
bad with Field.allInvertible F xCoeff1 xCoeffNonzero1
... | xinv , pr' = transitive (symmetric multIdentIsIdent) (transitive (multWellDefined (symmetric pr') reflexive) (transitive (symmetric multAssoc) (transitive (multWellDefined reflexive pr) (ringTimesZero R))))
LinearFunction.constant (composeLinearFunctions {_+_ = _+_} {_*_ = _*_} record { xCoeff = xCoeff1 ; xCoeffNonzero = xCoeffNonzero1 ; constant = constant1 } record { xCoeff = xCoeff2 ; xCoeffNonzero = xCoeffNonzero2 ; constant = constant2 }) = (xCoeff1 * constant2) + constant1
compositionIsCorrect : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A A A} {_*_ : A A A} {R : Ring S _+_ _*_} {F : Field R} (f1 : LinearFunction F) (f2 : LinearFunction F) {r : A} Setoid.__ S (interpretLinearFunction (composeLinearFunctions f1 f2) r) (((interpretLinearFunction f1) (interpretLinearFunction f2)) r)
compositionIsCorrect {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} record { xCoeff = xCoeff ; xCoeffNonzero = xCoeffNonzero ; constant = constant } record { xCoeff = xCoeff' ; xCoeffNonzero = xCoeffNonzero' ; constant = constant' } {r} = ans
where
open Setoid S
open Ring R
open Transitive (Equivalence.transitiveEq eq)
open Symmetric (Equivalence.symmetricEq eq)
open Reflexive (Equivalence.reflexiveEq eq)
ans : (((xCoeff * xCoeff') * r) + ((xCoeff * constant') + constant)) (xCoeff * ((xCoeff' * r) + constant')) + constant
ans = transitive (Group.multAssoc additiveGroup) (Group.wellDefined additiveGroup (transitive (Group.wellDefined additiveGroup (symmetric (Ring.multAssoc R)) reflexive) (symmetric (Ring.multDistributes R))) (reflexive {constant}))
linearFunctionsSetoid : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A A A} {_*_ : A A A} {R : Ring S _+_ _*_} (I : Field R) Setoid (LinearFunction I)
Setoid.__ (linearFunctionsSetoid {S = S} I) f1 f2 = ((LinearFunction.xCoeff f1) (LinearFunction.xCoeff f2)) && ((LinearFunction.constant f1) (LinearFunction.constant f2))
where
open Setoid S
Reflexive.reflexive (Equivalence.reflexiveEq (Setoid.eq (linearFunctionsSetoid {S = S} I))) = Reflexive.reflexive (Equivalence.reflexiveEq (Setoid.eq S)) ,, Reflexive.reflexive (Equivalence.reflexiveEq (Setoid.eq S))
Symmetric.symmetric (Equivalence.symmetricEq (Setoid.eq (linearFunctionsSetoid {S = S} I))) (fst ,, snd) = Symmetric.symmetric (Equivalence.symmetricEq (Setoid.eq S)) fst ,, Symmetric.symmetric (Equivalence.symmetricEq (Setoid.eq S)) snd
Transitive.transitive (Equivalence.transitiveEq (Setoid.eq (linearFunctionsSetoid {S = S} I))) (fst1 ,, snd1) (fst2 ,, snd2) = Transitive.transitive (Equivalence.transitiveEq (Setoid.eq S)) fst1 fst2 ,, Transitive.transitive (Equivalence.transitiveEq (Setoid.eq S)) snd1 snd2
linearFunctionsGroup : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A A A} {_*_ : A A A} {R : Ring S _+_ _*_} (F : Field R) Group (linearFunctionsSetoid F) (composeLinearFunctions)
Group.wellDefined (linearFunctionsGroup {R = R} F) {record { xCoeff = xCoeffM ; xCoeffNonzero = xCoeffNonzeroM ; constant = constantM }} {record { xCoeff = xCoeffN ; xCoeffNonzero = xCoeffNonzeroN ; constant = constantN }} {record { xCoeff = xCoeffX ; xCoeffNonzero = xCoeffNonzeroX ; constant = constantX }} {record { xCoeff = xCoeff ; xCoeffNonzero = xCoeffNonzero ; constant = constant }} (fst1 ,, snd1) (fst2 ,, snd2) = multWellDefined fst1 fst2 ,, Group.wellDefined additiveGroup (multWellDefined fst1 snd2) snd1
where
open Ring R
Group.identity (linearFunctionsGroup {S = S} {R = R} F) = record { xCoeff = Ring.1R R ; constant = Ring.0R R ; xCoeffNonzero = λ p Field.nontrivial F (Symmetric.symmetric (Equivalence.symmetricEq (Setoid.eq S)) p) }
Group.inverse (linearFunctionsGroup {S = S} {_*_ = _*_} {R = R} F) record { xCoeff = xCoeff ; constant = c ; xCoeffNonzero = pr } with Field.allInvertible F xCoeff pr
... | (inv , pr') = record { xCoeff = inv ; constant = inv * (Group.inverse (Ring.additiveGroup R) c) ; xCoeffNonzero = λ p Field.nontrivial F (transitive (symmetric (transitive (Ring.multWellDefined R p reflexive) (transitive (Ring.multCommutative R) (ringTimesZero R)))) pr') }
where
open Setoid S
open Transitive (Equivalence.transitiveEq (Setoid.eq S))
open Reflexive (Equivalence.reflexiveEq (Setoid.eq S))
open Symmetric (Equivalence.symmetricEq (Setoid.eq S))
Group.multAssoc (linearFunctionsGroup {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} F) {record { xCoeff = xA ; xCoeffNonzero = xANonzero ; constant = cA }} {record { xCoeff = xB ; xCoeffNonzero = xBNonzero ; constant = cB }} {record { xCoeff = xC ; xCoeffNonzero = xCNonzero ; constant = cC }} = Ring.multAssoc R ,, transitive (Group.wellDefined additiveGroup (transitive multDistributes (Group.wellDefined additiveGroup multAssoc reflexive)) reflexive) (symmetric (Group.multAssoc additiveGroup))
where
open Transitive (Equivalence.transitiveEq (Setoid.eq S))
open Reflexive (Equivalence.reflexiveEq (Setoid.eq S))
open Symmetric (Equivalence.symmetricEq (Setoid.eq S))
open Setoid S
open Ring R
Group.multIdentRight (linearFunctionsGroup {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} F) {record { xCoeff = xCoeff ; xCoeffNonzero = xCoeffNonzero ; constant = constant }} = transitive (Ring.multCommutative R) (Ring.multIdentIsIdent R) ,, transitive (Group.wellDefined additiveGroup (ringTimesZero R) reflexive) (Group.multIdentLeft additiveGroup)
where
open Ring R
open Setoid S
open Transitive (Equivalence.transitiveEq (Setoid.eq S))
open Reflexive (Equivalence.reflexiveEq (Setoid.eq S))
Group.multIdentLeft (linearFunctionsGroup {S = S} {R = R} F) {record { xCoeff = xCoeff ; xCoeffNonzero = xCoeffNonzero ; constant = constant }} = multIdentIsIdent ,, transitive (Group.multIdentRight additiveGroup) multIdentIsIdent
where
open Setoid S
open Ring R
open Transitive (Equivalence.transitiveEq (Setoid.eq S))
Group.invLeft (linearFunctionsGroup F) {record { xCoeff = xCoeff ; xCoeffNonzero = xCoeffNonzero ; constant = constant }} with Field.allInvertible F xCoeff xCoeffNonzero
Group.invLeft (linearFunctionsGroup {S = S} {R = R} F) {record { xCoeff = xCoeff ; xCoeffNonzero = xCoeffNonzero ; constant = constant }} | inv , prInv = prInv ,, transitive (symmetric multDistributes) (transitive (multWellDefined reflexive (Group.invRight additiveGroup)) (ringTimesZero R))
where
open Setoid S
open Ring R
open Transitive (Equivalence.transitiveEq (Setoid.eq S))
open Reflexive (Equivalence.reflexiveEq (Setoid.eq S))
open Symmetric (Equivalence.symmetricEq (Setoid.eq S))
Group.invRight (linearFunctionsGroup {S = S} {R = R} F) {record { xCoeff = xCoeff ; xCoeffNonzero = xCoeffNonzero ; constant = constant }} with Field.allInvertible F xCoeff xCoeffNonzero
... | inv , pr = transitive multCommutative pr ,, transitive (Group.wellDefined additiveGroup multAssoc reflexive) (transitive (Group.wellDefined additiveGroup (multWellDefined (transitive multCommutative pr) reflexive) reflexive) (transitive (Group.wellDefined additiveGroup multIdentIsIdent reflexive) (Group.invLeft additiveGroup)))
where
open Setoid S
open Ring R
open Transitive (Equivalence.transitiveEq (Setoid.eq S))
open Reflexive (Equivalence.reflexiveEq (Setoid.eq S))
open Symmetric (Equivalence.symmetricEq (Setoid.eq S))
{-
Question 3, part 2: prove that linearFunctionsGroup is not abelian
-}
-- We'll assume the field doesn't have characteristic 2.
linearFunctionsGroupNotAbelian : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A A A} {_*_ : A A A} {R : Ring S _+_ _*_} {F : Field R} (nonChar2 : Setoid.__ S ((Ring.1R R) + (Ring.1R R)) (Ring.0R R) False) AbelianGroup (linearFunctionsGroup F) False
linearFunctionsGroupNotAbelian {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} {F = F} pr record { commutative = commutative } = ans
where
open Ring R
open Group additiveGroup
open Symmetric (Equivalence.symmetricEq (Setoid.eq S)) renaming (symmetric to symmetricS)
open Transitive (Equivalence.transitiveEq (Setoid.eq S)) renaming (transitive to transitiveS)
open Reflexive (Equivalence.reflexiveEq (Setoid.eq S)) renaming (reflexive to reflexiveS)
f : LinearFunction F
f = record { xCoeff = 1R ; xCoeffNonzero = λ p Field.nontrivial F (symmetricS p) ; constant = 1R }
g : LinearFunction F
g = record { xCoeff = 1R + 1R ; xCoeffNonzero = pr ; constant = 0R }
gf : LinearFunction F
gf = record { xCoeff = 1R + 1R ; xCoeffNonzero = pr ; constant = 1R + 1R }
fg : LinearFunction F
fg = record { xCoeff = 1R + 1R ; xCoeffNonzero = pr ; constant = 1R }
oneWay : Setoid.__ (linearFunctionsSetoid F) gf (composeLinearFunctions g f)
oneWay = symmetricS (transitiveS multCommutative multIdentIsIdent) ,, transitiveS (symmetricS (transitiveS multCommutative multIdentIsIdent)) (symmetricS (Group.multIdentRight additiveGroup))
otherWay : Setoid.__ (linearFunctionsSetoid F) fg (composeLinearFunctions f g)
otherWay = symmetricS multIdentIsIdent ,, transitiveS (symmetricS (Group.multIdentLeft additiveGroup)) (Group.wellDefined additiveGroup (symmetricS multIdentIsIdent) (reflexiveS {1R}))
open Transitive (Equivalence.transitiveEq (Setoid.eq (linearFunctionsSetoid F)))
open Symmetric (Equivalence.symmetricEq (Setoid.eq (linearFunctionsSetoid F)))
bad : Setoid.__ (linearFunctionsSetoid F) gf fg
bad = transitive {gf} {composeLinearFunctions g f} {fg} oneWay (transitive {composeLinearFunctions g f} {composeLinearFunctions f g} {fg} (commutative {g} {f}) (symmetric {fg} {composeLinearFunctions f g} otherWay))
ans : False
ans with bad
ans | _ ,, contr = Field.nontrivial F (symmetricS (transitiveS {1R} {1R + (1R + Group.inverse additiveGroup 1R)} (transitiveS (symmetricS (Group.multIdentRight additiveGroup)) (Group.wellDefined additiveGroup reflexiveS (symmetricS (Group.invRight additiveGroup)))) (transitiveS (Group.multAssoc additiveGroup) (transitiveS (Group.wellDefined additiveGroup contr reflexiveS) (Group.invRight additiveGroup)))))

291
GroupsExamples.agda Normal file
View File

@@ -0,0 +1,291 @@
{-# OPTIONS --safe --warning=error #-}
open import Groups
open import Orders
open import Integers
open import Setoids
open import LogicalFormulae
open import FinSet
open import Functions
open import Naturals
open import IntegersModN
open import RingExamples
open import PrimeNumbers
open import Groups2
module GroupsExamples where
integersMinusNotGroup : Group (reflSetoid ) (_-Z_) False
integersMinusNotGroup record { wellDefined = wellDefined ; identity = identity ; inverse = inverse ; multAssoc = multAssoc ; multIdentRight = multIdentRight ; multIdentLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } with multAssoc {nonneg 3} {nonneg 2} {nonneg 1}
integersMinusNotGroup record { wellDefined = wellDefined ; identity = identity ; inverse = inverse ; multAssoc = multAssoc ; multIdentRight = multIdentRight ; multIdentLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } | ()
integersTimesNotGroup : Group (reflSetoid ) (_*Z_) False
integersTimesNotGroup record { wellDefined = wellDefined ; identity = (nonneg zero) ; inverse = inverse ; multAssoc = multAssoc ; multIdentRight = multIdentRight ; multIdentLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } with multIdentLeft {negSucc 1}
... | ()
integersTimesNotGroup record { wellDefined = wellDefined ; identity = (nonneg (succ zero)) ; inverse = inverse ; multAssoc = multAssoc ; multIdentRight = multIdentRight ; multIdentLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } with invLeft {nonneg zero}
... | bl with inverse (nonneg zero)
integersTimesNotGroup record { wellDefined = wellDefined ; identity = (nonneg (succ zero)) ; inverse = inverse ; multAssoc = multAssoc ; multIdentRight = multIdentRight ; multIdentLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } | () | nonneg zero
integersTimesNotGroup record { wellDefined = wellDefined ; identity = (nonneg (succ zero)) ; inverse = inverse ; multAssoc = multAssoc ; multIdentRight = multIdentRight ; multIdentLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } | () | nonneg (succ x)
integersTimesNotGroup record { wellDefined = wellDefined ; identity = (nonneg (succ zero)) ; inverse = inverse ; multAssoc = multAssoc ; multIdentRight = multIdentRight ; multIdentLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } | () | negSucc x
integersTimesNotGroup record { wellDefined = wellDefined ; identity = (nonneg (succ (succ x))) ; inverse = inverse ; multAssoc = multAssoc ; multIdentRight = multIdentRight ; multIdentLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } with succInjective (negSuccInjective (multIdentLeft {negSucc 1}))
... | ()
integersTimesNotGroup record { wellDefined = wellDefined ; identity = (negSucc x) ; inverse = inverse ; multAssoc = multAssoc ; multIdentRight = multIdentRight ; multIdentLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } with multIdentLeft {nonneg 2}
integersTimesNotGroup record { wellDefined = wellDefined ; identity = (negSucc x) ; inverse = inverse ; multAssoc = multAssoc ; multIdentRight = multIdentRight ; multIdentLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } | ()
-- Groups example 8.9 from lecture 1
data Weird : Set where
e : Weird
a : Weird
b : Weird
c : Weird
_+W_ : Weird Weird Weird
e +W t = t
a +W e = a
a +W a = e
a +W b = c
a +W c = b
b +W e = b
b +W a = c
b +W b = e
b +W c = a
c +W e = c
c +W a = b
c +W b = a
c +W c = e
+WAssoc : {x y z : Weird} (x +W (y +W z)) (x +W y) +W z
+WAssoc {e} {y} {z} = refl
+WAssoc {a} {e} {z} = refl
+WAssoc {a} {a} {e} = refl
+WAssoc {a} {a} {a} = refl
+WAssoc {a} {a} {b} = refl
+WAssoc {a} {a} {c} = refl
+WAssoc {a} {b} {e} = refl
+WAssoc {a} {b} {a} = refl
+WAssoc {a} {b} {b} = refl
+WAssoc {a} {b} {c} = refl
+WAssoc {a} {c} {e} = refl
+WAssoc {a} {c} {a} = refl
+WAssoc {a} {c} {b} = refl
+WAssoc {a} {c} {c} = refl
+WAssoc {b} {e} {z} = refl
+WAssoc {b} {a} {e} = refl
+WAssoc {b} {a} {a} = refl
+WAssoc {b} {a} {b} = refl
+WAssoc {b} {a} {c} = refl
+WAssoc {b} {b} {e} = refl
+WAssoc {b} {b} {a} = refl
+WAssoc {b} {b} {b} = refl
+WAssoc {b} {b} {c} = refl
+WAssoc {b} {c} {e} = refl
+WAssoc {b} {c} {a} = refl
+WAssoc {b} {c} {b} = refl
+WAssoc {b} {c} {c} = refl
+WAssoc {c} {e} {z} = refl
+WAssoc {c} {a} {e} = refl
+WAssoc {c} {a} {a} = refl
+WAssoc {c} {a} {b} = refl
+WAssoc {c} {a} {c} = refl
+WAssoc {c} {b} {e} = refl
+WAssoc {c} {b} {a} = refl
+WAssoc {c} {b} {b} = refl
+WAssoc {c} {b} {c} = refl
+WAssoc {c} {c} {e} = refl
+WAssoc {c} {c} {a} = refl
+WAssoc {c} {c} {b} = refl
+WAssoc {c} {c} {c} = refl
weirdGroup : Group (reflSetoid Weird) _+W_
Group.wellDefined weirdGroup = reflGroupWellDefined
Group.identity weirdGroup = e
Group.inverse weirdGroup t = t
Group.multAssoc weirdGroup {r} {s} {t} = +WAssoc {r} {s} {t}
Group.multIdentRight weirdGroup {e} = refl
Group.multIdentRight weirdGroup {a} = refl
Group.multIdentRight weirdGroup {b} = refl
Group.multIdentRight weirdGroup {c} = refl
Group.multIdentLeft weirdGroup {e} = refl
Group.multIdentLeft weirdGroup {a} = refl
Group.multIdentLeft weirdGroup {b} = refl
Group.multIdentLeft weirdGroup {c} = refl
Group.invLeft weirdGroup {e} = refl
Group.invLeft weirdGroup {a} = refl
Group.invLeft weirdGroup {b} = refl
Group.invLeft weirdGroup {c} = refl
Group.invRight weirdGroup {e} = refl
Group.invRight weirdGroup {a} = refl
Group.invRight weirdGroup {b} = refl
Group.invRight weirdGroup {c} = refl
weirdAb : AbelianGroup weirdGroup
AbelianGroup.commutative weirdAb {e} {e} = refl
AbelianGroup.commutative weirdAb {e} {a} = refl
AbelianGroup.commutative weirdAb {e} {b} = refl
AbelianGroup.commutative weirdAb {e} {c} = refl
AbelianGroup.commutative weirdAb {a} {e} = refl
AbelianGroup.commutative weirdAb {a} {a} = refl
AbelianGroup.commutative weirdAb {a} {b} = refl
AbelianGroup.commutative weirdAb {a} {c} = refl
AbelianGroup.commutative weirdAb {b} {e} = refl
AbelianGroup.commutative weirdAb {b} {a} = refl
AbelianGroup.commutative weirdAb {b} {b} = refl
AbelianGroup.commutative weirdAb {b} {c} = refl
AbelianGroup.commutative weirdAb {c} {e} = refl
AbelianGroup.commutative weirdAb {c} {a} = refl
AbelianGroup.commutative weirdAb {c} {b} = refl
AbelianGroup.commutative weirdAb {c} {c} = refl
weirdProjection : Weird FinSet 4
weirdProjection a = ofNat 0 (le 3 refl)
weirdProjection b = ofNat 1 (le 2 refl)
weirdProjection c = ofNat 2 (le 1 refl)
weirdProjection e = ofNat 3 (le zero refl)
weirdProjectionSurj : Surjection weirdProjection
Surjection.property weirdProjectionSurj fzero = a , refl
Surjection.property weirdProjectionSurj (fsucc fzero) = b , refl
Surjection.property weirdProjectionSurj (fsucc (fsucc fzero)) = c , refl
Surjection.property weirdProjectionSurj (fsucc (fsucc (fsucc fzero))) = e , refl
Surjection.property weirdProjectionSurj (fsucc (fsucc (fsucc (fsucc ()))))
weirdProjectionInj : (x y : Weird) weirdProjection x weirdProjection y Setoid.__ (reflSetoid Weird) x y
weirdProjectionInj e e fx=fy = refl
weirdProjectionInj e a ()
weirdProjectionInj e b ()
weirdProjectionInj e c ()
weirdProjectionInj a e ()
weirdProjectionInj a a fx=fy = refl
weirdProjectionInj a b ()
weirdProjectionInj a c ()
weirdProjectionInj b e ()
weirdProjectionInj b a ()
weirdProjectionInj b b fx=fy = refl
weirdProjectionInj b c ()
weirdProjectionInj c e ()
weirdProjectionInj c a ()
weirdProjectionInj c b ()
weirdProjectionInj c c fx=fy = refl
weirdFinite : FiniteGroup weirdGroup (FinSet 4)
SetoidToSet.project (FiniteGroup.toSet weirdFinite) = weirdProjection
SetoidToSet.wellDefined (FiniteGroup.toSet weirdFinite) x y = applyEquality weirdProjection
SetoidToSet.surj (FiniteGroup.toSet weirdFinite) = weirdProjectionSurj
SetoidToSet.inj (FiniteGroup.toSet weirdFinite) = weirdProjectionInj
FiniteSet.size (FiniteGroup.finite weirdFinite) = 4
FiniteSet.mapping (FiniteGroup.finite weirdFinite) = id
FiniteSet.bij (FiniteGroup.finite weirdFinite) = idIsBijective
weirdOrder : groupOrder weirdFinite 4
weirdOrder = refl
elementPowZ : (n : ) (elementPower Group (nonneg 1) n) nonneg n
elementPowZ zero = refl
elementPowZ (succ n) rewrite elementPowZ n = refl
Cyclic : CyclicGroup Group
CyclicGroup.generator Cyclic = nonneg 1
CyclicGroup.cyclic Cyclic {nonneg x} = inl (x , elementPowZ x)
CyclicGroup.cyclic Cyclic {negSucc x} = inr (succ x , ans)
where
ans : additiveInverseZ (nonneg 1 +Z elementPower Group (nonneg 1) x) negSucc x
ans rewrite elementPowZ x = refl
elementPowZn : (n : ) {pr : 0 <N succ (succ n)} (power : ) (powerLess : power <N succ (succ n)) {p : 1 <N succ (succ n)} elementPower (nGroup (succ (succ n)) pr) (record { x = 1 ; xLess = p }) power record { x = power ; xLess = powerLess }
elementPowZn n zero powerLess = equalityZn _ _ refl
elementPowZn n {pr} (succ power) powerLess {p} with orderIsTotal (n.x (elementPower (nGroup (succ (succ n)) pr) (record { x = 1 ; xLess = p }) power)) (succ n)
elementPowZn n {pr} (succ power) powerLess {p} | inl (inl x) = equalityZn _ _ (applyEquality succ v)
where
t : elementPower (nGroup (succ (succ n)) pr) (record { x = 1 ; xLess = succPreservesInequality (succIsPositive n) }) power record { x = power ; xLess = PartialOrder.transitive (TotalOrder.order TotalOrder) (a<SuccA power) powerLess }
t = elementPowZn n {pr} power (PartialOrder.transitive (TotalOrder.order TotalOrder) (a<SuccA power) powerLess) {succPreservesInequality (succIsPositive n)}
u : elementPower (nGroup (succ (succ n)) pr) (record { x = 1 ; xLess = p }) power record { x = power ; xLess = PartialOrder.transitive (TotalOrder.order TotalOrder) (a<SuccA power) powerLess }
u rewrite <NRefl p (succPreservesInequality (succIsPositive n)) = t
v : n.x (elementPower (nGroup (succ (succ n)) pr) (record { x = 1 ; xLess = p }) power) power
v rewrite u = refl
elementPowZn n {pr} (succ power) powerLess {p} | inl (inr x) with elementPower (nGroup (succ (succ n)) pr) (record { x = 1 ; xLess = p }) power
elementPowZn n {pr} (succ power) powerLess {p} | inl (inr x) | record { x = x' ; xLess = xLess } = exFalso (noIntegersBetweenXAndSuccX _ x xLess)
elementPowZn n {pr} (succ power) powerLess {p} | inr x with inspect (elementPower (nGroup (succ (succ n)) pr) (record { x = 1 ; xLess = p }) power)
elementPowZn n {pr} (succ power) powerLess {p} | inr x | record { x = x' ; xLess = p' } with inspected rewrite elementPowZn n {pr} power (PartialOrder.transitive (TotalOrder.order TotalOrder) (a<SuccA power) powerLess) {p} | x = exFalso (PartialOrder.irreflexive (TotalOrder.order TotalOrder) powerLess)
nCyclic : (n : ) (pr : 0 <N n) CyclicGroup (nGroup n pr)
nCyclic zero ()
CyclicGroup.generator (nCyclic (succ zero) pr) = record { x = 0 ; xLess = pr }
CyclicGroup.cyclic (nCyclic (succ zero) pr) {record { x = zero ; xLess = xLess }} rewrite <NRefl xLess (succIsPositive 0) | <NRefl pr (succIsPositive 0) = inl (0 , refl)
CyclicGroup.cyclic (nCyclic (succ zero) _) {record { x = succ x ; xLess = xLess }} = exFalso (zeroNeverGreater (canRemoveSuccFrom<N xLess))
nCyclic (succ (succ n)) pr = record { generator = record { x = 1 ; xLess = succPreservesInequality (succIsPositive n) } ; cyclic = λ {x} inl (ans x) }
where
ans : (z : n (succ (succ n)) pr) Sg (λ i (elementPower (nGroup (succ (succ n)) pr) (record { x = 1 ; xLess = succPreservesInequality (succIsPositive n) }) i) z)
ans record { x = x ; xLess = xLess } = x , elementPowZn n x xLess
multiplyByNHom : (n : ) GroupHom Group Group (λ i i *Z nonneg n)
GroupHom.groupHom (multiplyByNHom n) {x} {y} = lemma1
where
open Setoid (reflSetoid )
open Group Group
lemma : nonneg n *Z (x +Z y) (nonneg n) *Z x +Z nonneg n *Z y
lemma = additionZDistributiveMulZ (nonneg n) x y
lemma1 : (x +Z y) *Z nonneg n x *Z nonneg n +Z y *Z nonneg n
lemma1 rewrite multiplicationZIsCommutative (x +Z y) (nonneg n) | multiplicationZIsCommutative x (nonneg n) | multiplicationZIsCommutative y (nonneg n) = lemma
GroupHom.wellDefined (multiplyByNHom n) {x} {.x} refl = refl
modNExampleGroupHom : (n : ) (pr : 0 <N n) GroupHom Group (nGroup n pr) (mod n pr)
modNExampleGroupHom = {!!}
intoZn : {n : } {pr : 0 <N n} (x : ) (x<n : x <N n) mod n pr (nonneg x) record { x = x ; xLess = x<n }
intoZn {zero} {()}
intoZn {succ n} {0<n} x x<n with divisionAlg (succ n) x
intoZn {succ n} {0<n} x x<n | record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = inl y ; quotSmall = quotSmall } = equalityZn _ _ (modIsUnique (record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = inl y ; quotSmall = quotSmall }) (record { quot = 0 ; rem = x ; pr = ans ; remIsSmall = inl x<n ; quotSmall = inl (succIsPositive n)}))
where
ans : n *N 0 +N x x
ans rewrite multiplicationNIsCommutative n 0 = refl
intoZn {succ n} {0<n} x x<n | record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = inr () ; quotSmall = quotSmall }
quotientZn : (n : ) (pr : 0 <N n) GroupIso (quotientGroup Group (modNExampleGroupHom n pr)) (nGroup n pr) (mod n pr)
GroupHom.groupHom (GroupIso.groupHom (quotientZn n pr)) = GroupHom.groupHom (modNExampleGroupHom n pr)
GroupHom.wellDefined (GroupIso.groupHom (quotientZn n pr)) {x} {y} x~y = need
where
given : mod n pr (x +Z (Group.inverse Group y)) record { x = 0 ; xLess = pr }
given = x~y
given' : (mod n pr x) +n (mod n pr (Group.inverse Group y)) record { x = 0 ; xLess = pr }
given' = transitivity (equalityCommutative (GroupHom.groupHom (modNExampleGroupHom n pr))) given
given'' : (mod n pr x) +n Group.inverse (nGroup n pr) (mod n pr y) record { x = 0 ; xLess = pr}
given'' = transitivity (applyEquality (λ i mod n pr x +n i) (equalityCommutative (homRespectsInverse (modNExampleGroupHom n pr)))) given'
need : mod n pr x mod n pr y
need = transferToRight (nGroup n pr) given''
GroupIso.inj (quotientZn n pr) {x} {y} fx~fy = need
where
given : mod n pr x +n Group.inverse (nGroup n pr) (mod n pr y) record { x = 0 ; xLess = pr }
given = transferToRight'' (nGroup n pr) fx~fy
given' : mod n pr (x +Z (Group.inverse Group y)) record { x = 0 ; xLess = pr }
given' = identityOfIndiscernablesLeft _ _ _ _≡_ given (equalityCommutative (transitivity (GroupHom.groupHom (modNExampleGroupHom n pr) {x}) (applyEquality (λ i mod n pr x +n i) (homRespectsInverse (modNExampleGroupHom n pr)))))
need' : (mod n pr x) +n (mod n pr (Group.inverse Group y)) record { x = 0 ; xLess = pr }
need' rewrite equalityCommutative (GroupHom.groupHom (modNExampleGroupHom n pr) {x} {Group.inverse Group y}) = given'
need : mod n pr (x +Z (Group.inverse Group y)) record { x = 0 ; xLess = pr}
need = identityOfIndiscernablesLeft _ _ _ _≡_ need' (equalityCommutative (GroupHom.groupHom (modNExampleGroupHom n pr)))
GroupIso.surj (quotientZn n pr) {record { x = x ; xLess = xLess }} = nonneg x , intoZn x xLess
trivialGroupHom : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A A A} (G : Group S _+_) GroupHom trivialGroup G (λ _ Group.identity G)
GroupHom.groupHom (trivialGroupHom {S = S} G) = symmetric multIdentRight
where
open Setoid S
open Group G
open Symmetric (Equivalence.symmetricEq eq)
GroupHom.wellDefined (trivialGroupHom {S = S} G) _ = Reflexive.reflexive (Equivalence.reflexiveEq (Setoid.eq S))
trivialSubgroupIsNormal : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+A_ : A A A} (G : Group S _+A_) NormalSubgroup G trivialGroup (trivialGroupHom G)
Subgroup.fInj (NormalSubgroup.subgroup (trivialSubgroupIsNormal G)) {fzero} {fzero} _ = refl
Subgroup.fInj (NormalSubgroup.subgroup (trivialSubgroupIsNormal G)) {fzero} {fsucc ()} _
Subgroup.fInj (NormalSubgroup.subgroup (trivialSubgroupIsNormal G)) {fsucc ()} {y} _
NormalSubgroup.normal (trivialSubgroupIsNormal {S = S} {_+A_ = _+A_} G) = fzero , transitive (wellDefined (multIdentRight) reflexive) invRight
where
open Setoid S
open Group G
open Transitive (Equivalence.transitiveEq eq)
open Reflexive (Equivalence.reflexiveEq eq)
improperSubgroupIsNormal : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+A_ : A A A} (G : Group S _+A_) NormalSubgroup G G (identityHom G)
Subgroup.fInj (NormalSubgroup.subgroup (improperSubgroupIsNormal G)) = id
NormalSubgroup.normal (improperSubgroupIsNormal {S = S} {_+A_ = _+A_} G) {g} {h} = ((g +A h) +A inverse g) , reflexive
where
open Group G
open Setoid S
open Reflexive (Equivalence.reflexiveEq eq)

264
IntegerFactorisation.agda Normal file
View File

@@ -0,0 +1,264 @@
{-# OPTIONS --warning=error --safe #-}
open import LogicalFormulae
open import Naturals
open import PrimeNumbers
open import WellFoundedInduction
module IntegerFactorisation where
-- Represent a factorisation into increasing factors
-- Note that 0 cannot be expressed this way.
record factorisationNonunit (minFactor : ) (a : ) : Set where
inductive
field
1<a : 1 <N a
firstFactor :
firstFactorNontrivial : 1 <N firstFactor
firstFactorBiggerMin : minFactor ≤N firstFactor
firstFactorDivision : divisionAlgResult firstFactor a
firstFactorDivides : divisionAlgResult.rem firstFactorDivision 0
firstFactorPrime : Prime firstFactor
otherFactorsNumber :
otherFactors : ((divisionAlgResult.quot firstFactorDivision 1) && (otherFactorsNumber 0)) || (((1 <N divisionAlgResult.quot firstFactorDivision) && (factorisationNonunit firstFactor (divisionAlgResult.quot firstFactorDivision))))
lemma : (p : ) p *N 1 +N 0 p
lemma p rewrite addZeroRight (p *N 1) | productWithOneRight p = refl
lemma' : {a b : } a *N zero +N 0 b b zero
lemma' {a} {b} pr rewrite addZeroRight (a *N zero) | productZeroIsZeroRight a = equalityCommutative pr
primeFactorisation : {p : } (pr : Prime p) factorisationNonunit 1 p
primeFactorisation {p} record { p>1 = p>1 ; pr = pr } = record {1<a = p>1 ; firstFactor = p ; firstFactorNontrivial = p>1 ; firstFactorBiggerMin = inl p>1 ; firstFactorDivision = record { quot = 1 ; rem = 0 ; pr = lemma p ; remIsSmall = zeroIsValidRem p } ; firstFactorDivides = refl ; firstFactorPrime = record { p>1 = p>1 ; pr = pr} ; otherFactors = inl record { fst = refl ; snd = refl } ; otherFactorsNumber = 0 }
where
proof : (s : ) s *N 1 +N 0 s
proof s rewrite addZeroRight (s *N 1) | multiplicationNIsCommutative s 1 | addZeroRight s = refl
twoAsFact : factorisationNonunit 2 2
factorisationNonunit.1<a twoAsFact = succPreservesInequality (succIsPositive 0)
factorisationNonunit.firstFactor twoAsFact = 2
factorisationNonunit.firstFactorNontrivial twoAsFact = succPreservesInequality (succIsPositive 0)
factorisationNonunit.firstFactorBiggerMin twoAsFact = inr refl
factorisationNonunit.firstFactorDivision twoAsFact = record { quot = 1 ; rem = 0 ; remIsSmall = zeroIsValidRem 2 ; pr = refl}
factorisationNonunit.firstFactorDivides twoAsFact = refl
factorisationNonunit.firstFactorPrime twoAsFact = twoIsPrime
factorisationNonunit.otherFactorsNumber twoAsFact = 0
factorisationNonunit.otherFactors twoAsFact = inl record { fst = refl ; snd = refl }
fourFact : factorisationNonunit 1 4
factorisationNonunit.1<a fourFact = succPreservesInequality (succIsPositive 2)
factorisationNonunit.firstFactor fourFact = 2
factorisationNonunit.firstFactorNontrivial fourFact = succPreservesInequality (succIsPositive 0)
factorisationNonunit.firstFactorBiggerMin fourFact = inl (succPreservesInequality (succIsPositive 0))
factorisationNonunit.firstFactorDivision fourFact = record { quot = 2 ; rem = 0 ; remIsSmall = zeroIsValidRem 2 ; pr = refl}
factorisationNonunit.firstFactorDivides fourFact = refl
factorisationNonunit.firstFactorPrime fourFact = twoIsPrime
factorisationNonunit.otherFactorsNumber fourFact = 1
factorisationNonunit.otherFactors fourFact = inr record { fst = succPreservesInequality (succIsPositive 0) ; snd = twoAsFact }
lessLemma : {y : } (1 <N y) (zero <N y)
lessLemma {.(succ (x +N 1))} (le x refl) = succIsPositive (x +N 1)
canReduceFactorBound : {a i j : } factorisationNonunit i a j <N i factorisationNonunit j a
canReduceFactorBound {a} {i} {j} record { 1<a = 1<a ; firstFactor = firstFactor ; firstFactorNontrivial = firstFactorNontrivial ; firstFactorBiggerMin = inl ff<i ; firstFactorDivision = firstFactorDivision ; firstFactorDivides = firstFactorDivides ; firstFactorPrime = firstFactorPrime ; otherFactors = otherFactors } j<i = record { 1<a = 1<a ; firstFactor = firstFactor ; firstFactorNontrivial = firstFactorNontrivial ; firstFactorBiggerMin = inl (lessTransitive j<i ff<i) ; firstFactorDivision = firstFactorDivision ; firstFactorDivides = firstFactorDivides ; firstFactorPrime = firstFactorPrime ; otherFactors = otherFactors }
canReduceFactorBound {a} {i} {j} record { 1<a = 1<a ; firstFactor = firstFactor ; firstFactorNontrivial = firstFactorNontrivial ; firstFactorBiggerMin = inr ff=i ; firstFactorDivision = firstFactorDivision ; firstFactorDivides = firstFactorDivides ; firstFactorPrime = firstFactorPrime ; otherFactors = otherFactors } j<i = record { 1<a = 1<a ; firstFactor = firstFactor ; firstFactorNontrivial = firstFactorNontrivial ; firstFactorBiggerMin = inl (identityOfIndiscernablesRight j i firstFactor _<N_ j<i ff=i) ; firstFactorDivision = firstFactorDivision ; firstFactorDivides = firstFactorDivides ; firstFactorPrime = firstFactorPrime ; otherFactors = otherFactors }
canReduceFactorBound' : {a i j : } factorisationNonunit i a j ≤N i factorisationNonunit j a
canReduceFactorBound' {a} {i} {j} factA (inl x) = canReduceFactorBound factA x
canReduceFactorBound' {a} {i} {.i} factA (inr refl) = factA
canIncreaseFactorBound : {a i : } (fact : factorisationNonunit 1 a) ( x 1 <N x x <N i x a False) (i ≤N factorisationNonunit.firstFactor fact) factorisationNonunit i a
canIncreaseFactorBound {a} {i} record { 1<a = 1<a ; firstFactor = firstFactor ; firstFactorNontrivial = firstFactorNontrivial ; firstFactorBiggerMin = firstFactorBiggerMin ; firstFactorDivision = firstFactorDivision ; firstFactorDivides = firstFactorDivides ; firstFactorPrime = firstFactorPrime ; otherFactors = otherFactors } noSmaller iSmallEnough = record { 1<a = 1<a ; firstFactor = firstFactor ; firstFactorNontrivial = Prime.p>1 firstFactorPrime ; firstFactorBiggerMin = iSmallEnough ; firstFactorDivision = firstFactorDivision ; firstFactorDivides = firstFactorDivides ; firstFactorPrime = firstFactorPrime ; otherFactors = otherFactors }
-- Get the smallest prime factor of the number
everyNumberHasAPrimeFactor : {a : } (1 <N a) Sg (λ i ((i a) && (1 <N i)) && ((Prime i) && ( x x <N i 1 <N x x a False)))
everyNumberHasAPrimeFactor {a} 1<a with compositeOrPrime a 1<a
everyNumberHasAPrimeFactor {a} 1<a | inl record { n>1 = n>1 ; divisor = divisor ; dividesN = dividesN ; divisorLessN = divisorLessN ; divisorNot1 = divisorNot1 ; divisorPrime = divisorPrime ; noSmallerDivisors = noSmallerDivisors } = ( divisor , record { fst = record { fst = dividesN ; snd = divisorNot1 } ; snd = record { fst = divisorPrime ; snd = noSmallerDivisors } } )
everyNumberHasAPrimeFactor {a} 1<a | inr x = (a , record { fst = record { fst = aDivA a ; snd = 1<a } ; snd = record { fst = x ; snd = λ y y<a 1<y y|a lessImpliesNotEqual 1<y (equalityCommutative (Prime.pr x y|a y<a (lessLemma 1<y))) }} )
lemma2 : {a b c : } 1 <N a 0 <N b a *N b +N 0 c b <N c
lemma2 {zero} {b} {c} 1<a _ pr = exFalso (zeroNeverGreater 1<a)
lemma2 {succ zero} {b} {c} 1<a _ pr = exFalso (lessIrreflexive 1<a)
lemma2 {succ (succ a)} {zero} {zero} 1<a t pr = exFalso (lessIrreflexive t)
lemma2 {succ (succ a)} {zero} {succ c} 1<a t pr = succIsPositive c
lemma2 {succ (succ a)} {succ b} {c} 1<a t pr = le (b +N (a *N succ b)) go
where
assocLemm : (a b c : ) (a +N b) +N c (a +N c) +N b
assocLemm a b c rewrite additionNIsAssociative a b c | additionNIsCommutative b c | equalityCommutative (additionNIsAssociative a c b) = refl
go : succ ((b +N a *N succ b) +N succ b) c
go rewrite addZeroRight (succ (b +N succ (b +N a *N succ b))) | equalityCommutative (assocLemm b (succ b) (a *N succ b)) | additionNIsAssociative b (succ b) (a *N succ b) = pr
factorIntegerLemma : (x : ) (indHyp : (y : ) (y<x : y <N x) ((y <N 2) || (factorisationNonunit 1 y))) ((x <N 2) || (factorisationNonunit 1 x))
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 addZeroRight (a *N zero) | multiplicationNIsCommutative a 0 = 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} ; 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} ; 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') }
where
indHypRes : ((succ (succ qu)) <N 2) || factorisationNonunit 1 (succ (succ qu))
indHypRes = indHyp (succ (succ qu)) (lemma2 {a} {succ (succ qu)} {succ (succ x)} 1<a (succIsPositive (succ qu)) ssxDivA)
indHypRes' : factorisationNonunit 1 (succ (succ qu))
indHypRes' with indHypRes
indHypRes' | inl y = exFalso (zeroNeverGreater (canRemoveSuccFrom<N (canRemoveSuccFrom<N y)))
indHypRes' | inr y = y
z|ssx : (z : ) z succ (succ qu) z succ (succ x)
z|ssx z z|ssq = (divisibilityTransitive z|ssq (divides (record { quot = a ; rem = 0 ; pr = identityOfIndiscernablesLeft (a *N succ (succ qu) +N 0) (succ (succ x)) (succ (succ qu) *N a +N 0) _≡_ ssxDivA (applyEquality (λ t t +N 0) (multiplicationNIsCommutative a (succ (succ qu)))) ; remIsSmall = zeroIsValidRem (succ (succ qu))}) refl))
factNonunit : factorisationNonunit a (succ (succ qu))
factNonunit with orderIsTotal a (factorisationNonunit.firstFactor indHypRes')
factNonunit | inl (inl a<ff) = canIncreaseFactorBound indHypRes' (λ z 1<z z<a z|ssq smallerFactors z z<a 1<z (z|ssx z z|ssq)) (inl a<ff)
factNonunit | inl (inr ff<a) = exFalso (smallerFactors (factorisationNonunit.firstFactor indHypRes') ff<a (factorisationNonunit.firstFactorNontrivial indHypRes') (z|ssx (factorisationNonunit.firstFactor indHypRes') (divides (factorisationNonunit.firstFactorDivision indHypRes') (factorisationNonunit.firstFactorDivides indHypRes'))))
factNonunit | inr ff=a = canIncreaseFactorBound indHypRes' (λ z 1<z z<a z|ssq smallerFactors z z<a 1<z (divisibilityTransitive z|ssq (divides (record { quot = a ; rem = 0 ; pr = identityOfIndiscernablesLeft (a *N succ (succ qu) +N 0) (succ (succ x)) (succ (succ qu) *N a +N 0) _≡_ ssxDivA (applyEquality (λ t t +N 0) (multiplicationNIsCommutative a (succ (succ qu)))) ; remIsSmall = zeroIsValidRem (succ (succ qu))}) refl))) (inr ff=a)
factorInteger : (a : ) (1 <N a) factorisationNonunit 1 a
factorInteger a 1<a with (rec <NWellfounded (λ n (n <N 2) || (factorisationNonunit 1 n))) factorIntegerLemma
... | bl with bl a
factorInteger a 1<a | bl | inl x = exFalso (noIntegersBetweenXAndSuccX 1 1<a x)
factorInteger a 1<a | bl | inr x = x
lessTransLemma : {p i j : } p <N i i ≤N j p <N j
lessTransLemma {p} {i} {j} p<i (inl x) = orderIsTransitive p<i x
lessTransLemma {p} {i} {j} p<i (inr x) rewrite x = p<i
lemma4' : {quot rem b : } (quot +N quot) +N rem succ b quot <N succ b
lemma4' {zero} {rem} {b} pr = succIsPositive b
lemma4' {succ quot} {rem} {b} pr rewrite equalityCommutative (succExtracts (quot +N succ quot) rem) | additionNIsCommutative quot (succ quot) | additionNIsAssociative (succ quot) quot (succ rem) | succExtracts quot rem | additionNIsCommutative (succ quot) (succ (quot +N rem)) = le (quot +N rem) pr
lemma4 : {quot a rem b : } (a *N quot +N rem succ b) (1 <N a) (quot <N succ b)
lemma4 {quot} {zero} {rem} {b} pr 1<a = exFalso (zeroNeverGreater 1<a)
lemma4 {quot} {succ zero} {rem} {b} pr 1<a = exFalso (lessIrreflexive 1<a)
lemma4 {quot} {succ (succ zero)} {rem} {b} pr 1<a rewrite addZeroRight quot = lemma4' pr
lemma4 {quot} {succ (succ (succ a))} {rem} {b} pr 1<a = lemma4 {quot} {succ (succ a)} {quot +N rem} {b} p' (succPreservesInequality (succIsPositive a))
where
p' : (quot +N (quot +N a *N quot)) +N (quot +N rem) succ b
p' rewrite additionNIsCommutative quot (quot +N (quot +N a *N quot)) | additionNIsAssociative (quot +N (quot +N a *N quot)) quot rem = pr
noSmallerFactors : {a i p : } (factorisationNonunit i a) (Prime p) (p <N i) (p a) False
noSmallerFactors {a} {i} {p} factA pPrime p<i p|a with rec <NWellfounded (λ b (factorisationNonunit i b) p b False)
... | 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
where
x=firstFact : firstFactor *N 1 +N 0 x
x=firstFact rewrite equalityCommutative firstFactorDivides | equalityCommutative quotFirstfact=1 = divisionAlgResult.pr firstFactorDivision
x=firstFact' : firstFactor x
x=firstFact' = transitivity (equalityCommutative (lemma firstFactor)) x=firstFact
p|firstFact : p firstFactor
p|firstFact rewrite x=firstFact' = p|x
p=firstFact : p firstFactor
p=firstFact = primeDivPrimeImpliesEqual pPrime firstFactorPrime p|firstFact
i<=firstFactor : i ≤N p
i<=firstFactor rewrite p=firstFact = firstFactorBiggerMin
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
succXNotSmaller : succ x <N firstFactor False
succXNotSmaller = divisorIsSmaller {firstFactor} {x} (divides firstFactorDivision firstFactorDivides)
succXNotSmaller' : firstFactor ≤N succ x
succXNotSmaller' = notSmallerMeansGE succXNotSmaller
inter : firstFactor *N (divisionAlgResult.quot firstFactorDivision) +N divisionAlgResult.rem firstFactorDivision (succ x)
inter = divisionAlgResult.pr firstFactorDivision
inter' : firstFactor *N (divisionAlgResult.quot firstFactorDivision) +N 0 (succ x)
inter' rewrite equalityCommutative firstFactorDivides = inter
inter'' : firstFactor *N (divisionAlgResult.quot firstFactorDivision) (succ x)
inter'' rewrite equalityCommutative (addZeroRight (firstFactor *N (divisionAlgResult.quot firstFactorDivision))) = inter'
p|ff*q : p firstFactor *N (divisionAlgResult.quot firstFactorDivision)
p|ff*q rewrite inter'' = p|x
p|ffOrQ : (p firstFactor) || (p divisionAlgResult.quot firstFactorDivision)
p|ffOrQ = primesArePrime pPrime p|ff*q
p|ffIsFalse : (p firstFactor) False
p|ffIsFalse p|ff = lessIrreflexive (lessTransLemma p<i i<=p)
where
p=ff : p firstFactor
p=ff = primeDivPrimeImpliesEqual pPrime firstFactorPrime p|ff
i<=p : i ≤N p
i<=p rewrite p=ff = firstFactorBiggerMin
p|q : ((p firstFactor) || (p divisionAlgResult.quot firstFactorDivision)) (p divisionAlgResult.quot firstFactorDivision)
p|q (inl fls) = exFalso (p|ffIsFalse fls)
p|q (inr res) = res
lemma3 : {a : } a 0 1 <N a False
lemma3 {a} a=0 pr rewrite a=0 = zeroNeverGreater pr
firstFactorUniqueLemma : {i : } (a : ) (1 <N a) (f1 : factorisationNonunit i a) (f2 : factorisationNonunit i a) (factorisationNonunit.firstFactor f1 <N factorisationNonunit.firstFactor f2) False
firstFactorUniqueLemma {i} a 1<a f1 f2 ff1<ff2 = go
where
p1 = factorisationNonunit.firstFactor f1
rem1 = divisionAlgResult.quot (factorisationNonunit.firstFactorDivision f1)
p2 = factorisationNonunit.firstFactor f2
rem2 = divisionAlgResult.quot (factorisationNonunit.firstFactorDivision f2)
p1<p2 : p1 <N p2
p1<p2 = ff1<ff2
a=p2rem2 : a p2 *N rem2
a=p2rem2 with divisionAlgResult.pr (factorisationNonunit.firstFactorDivision f2)
... | ff rewrite factorisationNonunit.firstFactorDivides f2 | addZeroRight (factorisationNonunit.firstFactor f2 *N divisionAlgResult.quot (factorisationNonunit.firstFactorDivision f2)) = equalityCommutative ff
p1|second : (p1 p2) || (p1 rem2)
p1|second = primesArePrime {p1} {p2} {rem2} (factorisationNonunit.firstFactorPrime f1) lem
where
lem : p1 (p2 *N rem2)
lem = identityOfIndiscernablesRight p1 a (p2 *N rem2) __ (divides (factorisationNonunit.firstFactorDivision f1) (factorisationNonunit.firstFactorDivides f1)) a=p2rem2
p1|second' : ((p1 p2) || (p1 rem2)) ((p1 p2) || (p1 rem2))
p1|second' (inl x) = inl (primeDivPrimeImpliesEqual (factorisationNonunit.firstFactorPrime f1) (factorisationNonunit.firstFactorPrime f2) x)
p1|second' (inr x) = inr x
p1|second'' : (p1 p2) || (p1 rem2)
p1|second'' = p1|second' p1|second
go : False
go with p1|second''
go | inl x = lessImpliesNotEqual ff1<ff2 x
go | inr x with factorisationNonunit.otherFactors f2
go | inr x | inl record { fst = rem2=1 ; snd = _ } rewrite rem2=1 = exFalso (oneIsNotPrime res)
where
1prime' : Prime p1 Prime 1
1prime' = applyEquality Prime (oneHasNoDivisors x)
res : Prime 1
res rewrite equalityCommutative 1prime' = (factorisationNonunit.firstFactorPrime f1)
go | inr x | inr p1|rem2 with factorisationNonunit.otherFactors f2
go | inr x | inr p1|rem2 | inl record { fst = rem2=1 ; snd = _ } rewrite rem2=1 = exFalso (oneIsNotPrime res)
where
1prime' : Prime p1 Prime 1
1prime' = applyEquality Prime (oneHasNoDivisors x)
res : Prime 1
res rewrite equalityCommutative 1prime' = (factorisationNonunit.firstFactorPrime f1)
go | inr x | inr p1|rem2 | inr factorRem2 = noSmallerFactors (_&&_.snd factorRem2) (factorisationNonunit.firstFactorPrime f1) p1<p2 x
firstFactorUnique : {i : } (a : ) (1 <N a) (f1 : factorisationNonunit i a) (f2 : factorisationNonunit i a) (factorisationNonunit.firstFactor f1 factorisationNonunit.firstFactor f2)
firstFactorUnique {i} a 1<a f1 f2 with orderIsTotal (factorisationNonunit.firstFactor f1) (factorisationNonunit.firstFactor f2)
firstFactorUnique {i} a 1<a f1 f2 | inl (inl f1<f2) = exFalso (firstFactorUniqueLemma a 1<a f1 f2 f1<f2)
firstFactorUnique {i} a 1<a f1 f2 | inl (inr f2<f1) = exFalso (firstFactorUniqueLemma a 1<a f2 f1 f2<f1)
firstFactorUnique {i} a 1<a f1 f2 | inr x = x
factorListLemma : {i : } (a : ) (1 <N a) (f1 : factorisationNonunit i a) (f2 : factorisationNonunit i a) (divisionAlgResult.quot (factorisationNonunit.firstFactorDivision f2)) (divisionAlgResult.quot (factorisationNonunit.firstFactorDivision f1))
factorListLemma {i} a 1<a f1 f2 with firstFactorUnique {i} a 1<a f1 f2
... | firstFactEqual = res
where
div1 : divisionAlgResult (factorisationNonunit.firstFactor f1) a
div1 = factorisationNonunit.firstFactorDivision f1
rem1=0 : divisionAlgResult.rem div1 0
rem1=0 = factorisationNonunit.firstFactorDivides f1
pr1 : (factorisationNonunit.firstFactor f1) *N (divisionAlgResult.quot div1) +N 0 a
pr1 rewrite equalityCommutative rem1=0 = divisionAlgResult.pr div1
pr1' : (factorisationNonunit.firstFactor f1) *N (divisionAlgResult.quot div1) a
pr1' rewrite equalityCommutative (addZeroRight ((factorisationNonunit.firstFactor f1) *N (divisionAlgResult.quot div1))) = pr1
div2 : divisionAlgResult (factorisationNonunit.firstFactor f2) a
div2 = factorisationNonunit.firstFactorDivision f2
rem2=0 : divisionAlgResult.rem div2 0
rem2=0 = factorisationNonunit.firstFactorDivides f2
pr2 : (factorisationNonunit.firstFactor f2) *N (divisionAlgResult.quot div2) +N 0 a
pr2 rewrite equalityCommutative rem2=0 = divisionAlgResult.pr div2
pr2' : (factorisationNonunit.firstFactor f2) *N (divisionAlgResult.quot div2) a
pr2' rewrite equalityCommutative (addZeroRight ((factorisationNonunit.firstFactor f2) *N (divisionAlgResult.quot div2))) = pr2
pr : (factorisationNonunit.firstFactor f2) *N (divisionAlgResult.quot div2) (factorisationNonunit.firstFactor f1) *N (divisionAlgResult.quot div1)
pr = transitivity pr2' (equalityCommutative pr1')
pr' : (factorisationNonunit.firstFactor f1) *N (divisionAlgResult.quot div2) (factorisationNonunit.firstFactor f1) *N (divisionAlgResult.quot div1)
pr' = identityOfIndiscernablesLeft ((factorisationNonunit.firstFactor f2) *N (divisionAlgResult.quot div2)) ((factorisationNonunit.firstFactor f1) *N (divisionAlgResult.quot div1)) ((factorisationNonunit.firstFactor f1) *N (divisionAlgResult.quot div2)) _≡_ pr (applyEquality (λ t t *N divisionAlgResult.quot div2) (equalityCommutative firstFactEqual))
positive : zero <N factorisationNonunit.firstFactor f1
positive = lessTransitive (succIsPositive 0) (factorisationNonunit.firstFactorNontrivial f1)
res : divisionAlgResult.quot div2 divisionAlgResult.quot div1
res = productCancelsLeft (factorisationNonunit.firstFactor f1) (divisionAlgResult.quot div2) (divisionAlgResult.quot div1) positive pr'
factorListSameLength : {i : } (a : ) (1 <N a) (f1 : factorisationNonunit i a) (f2 : factorisationNonunit i a) (divisionAlgResult.quot (factorisationNonunit.firstFactorDivision f1) 1) divisionAlgResult.quot (factorisationNonunit.firstFactorDivision f2) 1
factorListSameLength {i} a 1<a f1 f2 quot=1 with firstFactorUnique {i} a 1<a f1 f2
... | firstFactEqual with factorListLemma {i} a 1<a f1 f2
... | eq = transitivity eq quot=1

1541
Integers.agda Normal file

File diff suppressed because it is too large Load Diff

696
IntegersModN.agda Normal file
View File

@@ -0,0 +1,696 @@
{-# OPTIONS --safe --warning=error #-}
open import LogicalFormulae
open import Groups
open import Naturals
open import PrimeNumbers
open import Setoids
open import FinSet
open import Functions
module IntegersModN 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
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 equalityCommutative (additionNIsAssociative (succ x) a b) | additionNIsCommutative x a | additionNIsAssociative (succ a) x b | additionNIsCommutative a (x +N b) | additionNIsCommutative (succ (x +N b)) a = canSubtractFromEqualityLeft {a} proof
_+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 orderIsTotal (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) (a +n record { x = 0 ; xLess = pr }) a
plusZnIdentityRight {zero} {()} a
plusZnIdentityRight {succ x} {_} record { x = a ; xLess = aLess } with orderIsTotal (a +N 0) (succ x)
plusZnIdentityRight {succ x} {_} record { x = a ; xLess = aLess } | inl (inl a<sx) = equalityZn _ _ (additionNIsCommutative 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 additionNIsCommutative a 0 = orderIsIrreflexive 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 additionNIsCommutative a 0 | a=sx = lessIrreflexive aL
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 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 } | inr x=succn rewrite x=succn = exFalso (lessIrreflexive 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)))
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 orderIsTotal (a +N b) (succ n)
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inl (inl a+b<sn) with orderIsTotal (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 _ _ (additionNIsCommutative 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)
where
f : (a +N b) <N succ n succ n <N b +N a False
f pr1 pr2 rewrite additionNIsCommutative b a = lessIrreflexive (orderIsTransitive 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 additionNIsCommutative b a | eq = lessIrreflexive pr1
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inl (inr sn<a+b) with orderIsTotal (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 additionNIsCommutative b a = lessIrreflexive (orderIsTransitive 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 (additionNIsCommutative 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 additionNIsCommutative b a | eq = lessIrreflexive pr
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inr sn=a+b with orderIsTotal (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 additionNIsCommutative b a | sn=a+b = lessIrreflexive 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 additionNIsCommutative b a | sn=a+b = lessIrreflexive 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
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 = lessIrreflexive (orderIsTransitive 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) (additionNIsAssociative 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 additionNIsAssociative a b c = pr2
pr4 : n +N n <N c +N n
pr4 rewrite additionNIsCommutative c n = identityOfIndiscernablesRight _ _ _ _<N_ pr3 (applyEquality (_+N c) a+b=n)
pr5 : n <N c
pr5 = subtractionPreservesInequality n pr4
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) | equalityCommutative (additionNIsAssociative 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 equalityCommutative (additionNIsAssociative a b c) | additionNIsCommutative a b | additionNIsAssociative b a c = refl
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 = lessIrreflexive (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) (equalityCommutative (additionNIsAssociative 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 equalityCommutative (additionNIsAssociative 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
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 = lessIrreflexive (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) (additionNIsAssociative a _ n)
p2 : (a +N b) +N c n +N n
p2 rewrite additionNIsAssociative 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 equalityCommutative (additionNIsAssociative 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 additionNIsCommutative b 0 | a+b=sn = exFalso (lessIrreflexive a+b+c<sn)
help27 {a} {b} {succ c} {n} a+b=sn a+b+c<sn rewrite equalityCommutative (additionNIsAssociative 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 (equalityCommutative (additionNIsAssociative a b c)) | a+b=n | additionNIsCommutative 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 additionNIsCommutative a 0 | additionNIsCommutative 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 additionNIsCommutative a 0 = lessIrreflexive (orderIsTransitive 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 additionNIsCommutative a 0 | a+0=n = lessIrreflexive 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 = lessIrreflexive (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) (additionNIsAssociative 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 (equalityCommutative (additionNIsAssociative 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 = lessIrreflexive (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 | additionNIsCommutative 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) (succ c) b+c=c = b=0' (succ b) c bl2
where
bl2 : succ b +N c c
bl2 rewrite additionNIsCommutative b c | additionNIsCommutative (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 = lessIrreflexive (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) (additionNIsAssociative 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 additionNIsAssociative a b c = pr2
pr4 : c +N n <N n +N n
pr4 = identityOfIndiscernablesLeft _ _ _ _<N_ pr3 (transitivity (applyEquality (_+N c) a+b=n) (additionNIsCommutative 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 = lessIrreflexive (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) (equalityCommutative (additionNIsAssociative 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 (additionNIsAssociative 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 = lessIrreflexive (orderIsTransitive 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) (equalityCommutative (additionNIsAssociative 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' (additionNIsAssociative a b c)
p3 : n +N n <N a +N n
p3 = orderIsTransitive 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 = lessIrreflexive (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) (additionNIsAssociative 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) (equalityCommutative (additionNIsAssociative 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'' (additionNIsAssociative 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 (lessIrreflexive (orderIsTransitive 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) (additionNIsAssociative 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) (equalityCommutative (additionNIsAssociative 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'' (additionNIsAssociative 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 = lessIrreflexive (orderIsTransitive 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) (equalityCommutative (additionNIsAssociative 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) (additionNIsAssociative 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' (equalityCommutative (additionNIsAssociative 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 (equalityCommutative (additionNIsAssociative n _ c)) (transitivity (applyEquality (_+N c) (addMinus' (inl n<a+b))) (transitivity (additionNIsAssociative 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 = additionNIsAssociative 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 = lessIrreflexive (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) (additionNIsAssociative 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) (equalityCommutative (additionNIsAssociative 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 equalityCommutative (additionNIsAssociative 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 = lessIrreflexive lemm4
where
pr : {a b c : } a +N (b +N c) b +N (a +N c)
pr {a} {b} {c} rewrite equalityCommutative (additionNIsAssociative a b c) | additionNIsCommutative a b | additionNIsAssociative 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) (equalityCommutative (additionNIsAssociative 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) (additionNIsAssociative 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 additionNIsAssociative 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 = lessIrreflexive (orderIsTransitive a<n lemm5)
where
pr : {a b c : } a +N (b +N c) b +N (a +N c)
pr {a} {b} {c} rewrite equalityCommutative (additionNIsAssociative a b c) | additionNIsCommutative a b | additionNIsAssociative 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) (equalityCommutative (additionNIsAssociative 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 (additionNIsAssociative 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 = lessIrreflexive lemm6
where
pr : {a b c : } a +N (b +N c) b +N (a +N c)
pr {a} {b} {c} rewrite equalityCommutative (additionNIsAssociative a b c) | additionNIsCommutative a b | additionNIsAssociative 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) (equalityCommutative (additionNIsAssociative 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 (additionNIsAssociative 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 additionNIsCommutative a 0 | n=a+0 = lessIrreflexive a<n
help8 : {a n : } (n <N a +N 0) (a <N n) False
help8 {a} {n} n<a+0 a<n rewrite additionNIsCommutative a 0 = lessIrreflexive (orderIsTransitive 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 additionNIsCommutative 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) | additionNIsAssociative a b c | b+c=n = additionNIsCommutative n a
lem' : n +N a n +N (subtractionNResult.result (-N (inl n<a+b)) +N c)
lem' = identityOfIndiscernablesRight _ _ _ _≡_ lem (additionNIsAssociative 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) = equalityCommutative (additionNIsAssociative 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 (additionNIsAssociative 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 equalityCommutative (additionNIsAssociative a b c) | additionNIsCommutative a b | additionNIsAssociative 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) = equalityCommutative (additionNIsAssociative 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 (additionNIsAssociative 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 = lessIrreflexive (orderIsTransitive (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 equalityCommutative (additionNIsAssociative a b c) | equalityCommutative (additionNIsAssociative b a c) = applyEquality (_+N c) (additionNIsCommutative 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 | equalityCommutative (additionNIsAssociative a b c) = additionPreservesInequality c a+b<n
inter4 : (n +N n <N n +N c) n <N c
inter4 pr rewrite additionNIsCommutative 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) = equalityCommutative (additionNIsAssociative 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 (equalityCommutative (additionNIsAssociative 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 | additionNIsCommutative n 0 = lessIrreflexive 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 equalityCommutative (additionNIsAssociative a (succ n) b+c-sn) | additionNIsCommutative a (succ n) | additionNIsAssociative (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 equalityCommutative (additionNIsAssociative a b c) = additionPreservesInequality c a+b<sn
eep2' : a +N (b +N c) <N succ n +N succ n
eep2' = orderIsTransitive eep2 (additionPreservesInequalityOnLeft (succ n) c<sn)
ans : False
ans = lessIrreflexive (orderIsTransitive 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 orderIsTotal (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 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 (inl a+b<sn) | inl (inl a+b+c<sn) 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 (inl a+b<sn) | inl (inl a+b+c<sn) | 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 (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inl b+c<sn) | inl (inl a+'b+c<sn) = equalityZn _ _ (equalityCommutative (additionNIsAssociative 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 additionNIsAssociative a b c = lessIrreflexive (orderIsTransitive 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 additionNIsAssociative a b c | p2 = lessIrreflexive 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 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 (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inr sn<b+c) | inl (inl _) with orderIsTotal (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 | equalityCommutative (additionNIsAssociative a (succ n) result) | additionNIsCommutative a (succ n) | additionNIsAssociative (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 (equalityCommutative (additionNIsAssociative a (succ n) result)) (transitivity (applyEquality (λ t t +N result) (additionNIsCommutative a (succ n))) (additionNIsAssociative (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 (equalityCommutative (additionNIsAssociative a (succ n) result)) (transitivity (applyEquality (λ t t +N result) (additionNIsCommutative a (succ n))) (additionNIsAssociative (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 additionNIsAssociative a b c = lessIrreflexive (orderIsTransitive 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 | additionNIsAssociative a b c = lessIrreflexive 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 additionNIsAssociative a b c | pr2 | additionNIsCommutative 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 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 (inl a+b<sn) | inl (inr sn<a+b+c) | 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 (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 additionNIsAssociative a b c = lessIrreflexive (orderIsTransitive 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 additionNIsAssociative a b c = additionNIsCommutative (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 additionNIsAssociative a b c | pr2 = lessIrreflexive 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 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 (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 additionNIsAssociative a b c = lessIrreflexive (orderIsTransitive 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 orderIsTotal (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 additionNIsAssociative a b c | pr2 = lessIrreflexive 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 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 (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 additionNIsAssociative a b c = lessIrreflexive (orderIsTransitive 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 orderIsTotal (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 additionNIsCommutative a 0 | additionNIsAssociative a b c | b+c=sn | additionNIsCommutative (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 additionNIsCommutative a 0 = lessIrreflexive (orderIsTransitive pr1 pr2)
false (succ b) pr1 pr2 rewrite additionNIsCommutative a 0 = lessIrreflexive (orderIsTransitive pr2 (orderIsTransitive (le b (additionNIsCommutative (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 = lessIrreflexive pr1
false (succ b) pr1 pr2 rewrite additionNIsCommutative 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 additionNIsAssociative a b c | pr2 = lessIrreflexive 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 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 (inl a+b<sn) | inr sn=a+b+c | 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 (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 additionNIsAssociative a b c | pr1 = lessIrreflexive 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 additionNIsAssociative a b c | pr1 = lessIrreflexive 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 additionNIsAssociative a b c | equalityCommutative pr1 = lessIrreflexive pr2
false (succ a) pr1 pr2 rewrite additionNIsAssociative a b c | equalityCommutative pr1 = lessIrreflexive (orderIsTransitive 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 orderIsTotal (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 additionNIsAssociative a b c | pr2 = refl
a=0 (succ a) pr1 pr2 rewrite additionNIsAssociative a b c | pr2 = canSubtractFromEqualityRight pr1
ans : a +N 0 0
ans rewrite additionNIsCommutative 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 additionNIsAssociative a b c | pr3 | additionNIsCommutative 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 additionNIsAssociative a b c | pr2 | equalityCommutative pr3 | additionNIsCommutative 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 = 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)
where
false : succ n <N a +N b a +N (b +N c) <N succ n False
false pr1 pr2 rewrite equalityCommutative (additionNIsAssociative a b c) = cannotAddAndEnlarge' (orderIsTransitive 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 orderIsTotal (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 equalityCommutative (additionNIsAssociative 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 orderIsTotal (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 orderIsTotal (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 orderIsTotal (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 orderIsTotal (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 orderIsTotal (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 orderIsTotal (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 orderIsTotal 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 orderIsTotal (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 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 } | 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 orderIsTotal (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 orderIsTotal (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 (lessIrreflexive (orderIsTransitive 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 orderIsTotal (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 | additionNIsCommutative 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 orderIsTotal (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 (lessIrreflexive 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 (subtractOneFromInequality 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)
... | 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)
f : False
f rewrite h = lessIrreflexive x
... | inl (inr x) = exFalso f
where
h : subtractionNResult.result (-N (inl (subtractOneFromInequality aLess))) +N succ a succ n
h = addMinus {succ a} {succ n} (inl aLess)
f : False
f rewrite h = lessIrreflexive x
... | (inr n-a+sa=sn) = equalityZn _ _ refl
nGroup : (n : ) (pr : 0 <N n) Group (reflSetoid (n n pr)) _+n_
nGroup n pr = record { identity = record { x = 0 ; xLess = pr } ; inverse = λ a underlying (inverseZn a) ; multAssoc = λ {a} {b} {c} plusZnAssociative a b c ; multIdentRight = λ {a} plusZnIdentityRight a ; multIdentLeft = λ {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
nAbGroup : (n : ) (pr : 0 <N n) AbelianGroup (nGroup n pr)
AbelianGroup.commutative (nAbGroup n pr) {a} {b} = plusZnCommutative 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

66
IntegersModNRing.agda Normal file
View File

@@ -0,0 +1,66 @@
{-# OPTIONS --safe --warning=error #-}
open import LogicalFormulae
open import Orders
open import Groups
open import Naturals
open import PrimeNumbers
open import Rings
open import Setoids
open import IntegersModN
module IntegersModNRing where
modNTo : {n : } {pr : 0 <N n} (a : n n pr)
modNTo record { x = x ; xLess = xLess } = x
*nhelper : {n : } {pr : 0 <N n} (max : ) (a : n n pr) n n pr (max modNTo a) n n pr
*nhelper {zero} {()}
*nhelper {succ n} zero a b pr2 = record { x = 0 ; xLess = succIsPositive n }
*nhelper {succ n} (succ max) record { x = zero ; xLess = xLess } b ()
*nhelper {succ n} (succ max) record { x = (succ x) ; xLess = xLess } b pr2 = (*nhelper max record { x = x ; xLess = xLess2 } b (succInjective pr2)) +n b
where
xLess2 : x <N succ n
xLess2 = PartialOrder.transitive (TotalOrder.order TotalOrder) (a<SuccA x) xLess
_*n_ : {n : } {pr : 0 <N n} n n pr n n pr n n pr
_*n_ {0} {()}
_*n_ {succ n} {_} record { x = a ; xLess = aLess } b = *nhelper a record { x = a ; xLess = aLess } b refl
nIdent : (n : ) (pr : 0 <N n) n n pr
nIdent 0 ()
nIdent 1 pr = record { x = 0 ; xLess = pr }
nIdent (succ (succ n)) _ = record { x = 1 ; xLess = succPreservesInequality (succIsPositive n) }
nMult0Right : {n : } {pr : 0 <N n} (max : ) (a : n n pr) (modNTo a max) (a *n record { x = 0 ; xLess = pr }) record { x = 0 ; xLess = pr }
nMult0Right {0} {()}
nMult0Right {succ n} .zero record { x = zero ; xLess = xLess } refl = equalityZn _ _ refl
nMult0Right {succ n} {pr} (.succ x) record { x = (succ x) ; xLess = xLess } refl rewrite nMult0Right {succ n} {pr} x record { x = x ; xLess = PartialOrder.transitive (TotalOrder.order TotalOrder) (a<SuccA x) xLess } refl = equalityZn _ _ refl
nMultCommutative : {n : } {pr : 0 <N n} (a b : n n pr) (a *n b) (b *n a)
nMultCommutative {0} {()}
nMultCommutative {succ n} {pr} record { x = zero ; xLess = xLess } b with <NRefl pr xLess
... | refl rewrite nMult0Right (modNTo b) b refl = equalityZn _ _ refl
nMultCommutative {succ n} record { x = (succ x) ; xLess = xLess } b = {!!}
nMultIdent : {n : } {pr : 0 <N n} (a : n n pr) (nIdent n pr) *n a a
nMultIdent {zero} {()}
nMultIdent {succ zero} {pr} record { x = zero ; xLess = (le diff proof) } = equalityZn _ _ refl
nMultIdent {succ zero} {pr} record { x = (succ a) ; xLess = (le diff proof) } = exFalso f
where
f : False
f rewrite additionNIsCommutative diff (succ a) = naughtE (succInjective (equalityCommutative proof))
nMultIdent {succ (succ n)} {pr} record { x = a ; xLess = aLess } with orderIsTotal a (succ (succ n))
nMultIdent {succ (succ n)} {pr} record { x = a ; xLess = aLess } | inl (inl a<ssn) = equalityZn _ _ refl
nMultIdent {succ (succ n)} {pr} record { x = a ; xLess = aLess } | inl (inr ssn<a) = exFalso (PartialOrder.irreflexive (TotalOrder.order TotalOrder) (PartialOrder.transitive (TotalOrder.order TotalOrder) aLess ssn<a))
nMultIdent {succ (succ n)} {pr} record { x = .(succ (succ n)) ; xLess = aLess } | inr refl = exFalso (PartialOrder.irreflexive (TotalOrder.order TotalOrder) aLess)
nRing : (n : ) (pr : 0 <N n) Ring {_} (n n pr)
Ring.additiveGroup (nRing n 0<n) = AbelianGroup.grp (nGroup n 0<n)
Ring._*_ (nRing n 0<n) a b = a *n b
Ring.multWellDefined (nRing n 0<n) = reflGroupWellDefined
Ring.1R (nRing n pr) = nIdent n pr
Ring.groupIsAbelian (nRing n 0<n) = AbelianGroup.commutative (nGroup n 0<n)
Ring.multAssoc (nRing n 0<n) = {!!}
Ring.multCommutative (nRing n 0<n) {a} {b} = nMultCommutative a b
Ring.multDistributes (nRing n 0<n) = {!!}
Ring.multIdentIsIdent (nRing n 0<n) {a} = nMultIdent a

17
IntegralDomains.agda Normal file
View File

@@ -0,0 +1,17 @@
{-# OPTIONS --safe --warning=error #-}
open import LogicalFormulae
open import Groups
open import Naturals
open import Orders
open import Setoids
open import Functions
open import Rings
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
module IntegralDomains where
record IntegralDomain {n m : _} {A : Set n} {S : Setoid {n} {m} A} {_+_ : A A A} {_*_ : A A A} (R : Ring S _+_ _*_) : Set (lsuc m n) where
field
intDom : {a b : A} Setoid.__ S (a * b) (Ring.0R R) (Setoid.__ S a (Ring.0R R)) || (Setoid.__ S b (Ring.0R R))
nontrivial : Setoid.__ S (Ring.1R R) (Ring.0R R) False

231
KeyValue.agda Normal file
View File

@@ -0,0 +1,231 @@
{-# OPTIONS --warning=error --safe #-}
open import LogicalFormulae
open import Orders
open import Maybe
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Vectors
open import Naturals
module KeyValue where
record ReducedMap {a b c : _} (keys : Set a) (values : Set b) (keyOrder : TotalOrder {_} {c} keys) (min : keys) : Set (a b c)
record ReducedMap {a} {b} {c} keys values keyOrder min where
inductive
field
firstEntry : values
next : Maybe (Sg keys (λ nextKey (ReducedMap keys values keyOrder nextKey) && (TotalOrder._<_ keyOrder min nextKey)))
addReducedMap : {a b c : _} {keys : Set a} {values : Set b} {keyOrder : TotalOrder {_} {c} keys} {min : keys} (k : keys) (v : values) (m : ReducedMap {a} {b} keys values keyOrder min) ReducedMap keys values keyOrder (TotalOrder.min keyOrder min k)
addReducedMap {keyOrder = keyOrder} {min} k v m with TotalOrder.totality keyOrder min k
addReducedMap {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = no } | inl (inl min<k) = record { firstEntry = firstEntry ; next = yes (k , (record { firstEntry = v ; next = no} ,, min<k))}
addReducedMap {keyOrder = keyOrder} {min} k v record { firstEntry = minVal ; next = yes (nextKey , (m ,, pr)) } | inl (inl min<k) = record { firstEntry = minVal ; next = yes ((TotalOrder.min keyOrder nextKey k) , (addReducedMap {keyOrder = keyOrder} {_} k v m ,, minFromBoth {order = keyOrder} pr min<k))}
addReducedMap {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = next } | inl (inr k<min) = record { firstEntry = v ; next = yes (min , (record { firstEntry = firstEntry ; next = next } ,, k<min)) }
addReducedMap {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = next } | inr min=k rewrite min=k = record { firstEntry = v ; next = next }
lookupReduced : {a b c : _} {keys : Set a} {values : Set b} {keyOrder : TotalOrder {_} {c} keys} {min : keys} (m : ReducedMap keys values keyOrder min) (target : keys) Maybe values
lookupReduced {keyOrder = keyOrder} {min} m k with TotalOrder.totality keyOrder min k
lookupReduced {keyOrder = keyOrder} {min} record { firstEntry = firstEntry ; next = no } k | inl (inl min<k) = no
lookupReduced {keyOrder = keyOrder} {min} record { firstEntry = firstEntry ; next = (yes (newMin , (m ,, _))) } k | inl (inl min<k) = lookupReduced {keyOrder = keyOrder} {newMin} m k
lookupReduced {keyOrder = keyOrder} {min} m k | inl (inr k<min) = no
lookupReduced {keyOrder = keyOrder} {min} record { firstEntry = firstEntry ; next = next } k | inr min=k = yes firstEntry
countReduced : {a b c : _} {keys : Set a} {values : Set b} {keyOrder : TotalOrder {_} {c} keys} {min : keys} (m : ReducedMap keys values keyOrder min)
countReduced record { firstEntry = firstEntry ; next = no } = 1
countReduced record { firstEntry = firstEntry ; next = (yes (key , (m ,, pr))) } = succ (countReduced m)
lookupReducedSucceedsAfterAdd : {a b c : _} {keys : Set a} {values : Set b} {keyOrder : TotalOrder {_} {c} keys} {min : keys} (k : keys) (v : values) (m : ReducedMap keys values keyOrder min) (lookupReduced (addReducedMap k v m) k yes v)
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v m with TotalOrder.totality keyOrder min k
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = no } | inl (inl min<k) with TotalOrder.totality keyOrder min k
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = no } | inl (inl min<k) | inl (inl _) with TotalOrder.totality keyOrder k k
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = no } | inl (inl min<k) | inl (inl _) | inl (inl k<k) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) k<k)
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = no } | inl (inl min<k) | inl (inl _) | inl (inr k<k) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) k<k)
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = no } | inl (inl min<k) | inl (inl _) | inr refl = refl
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = no } | inl (inl min<k) | inl (inr k<min) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) (PartialOrder.transitive (TotalOrder.order keyOrder) min<k k<min))
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = no } | inl (inl min<k) | inr min=k rewrite min=k = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) min<k)
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = (yes (newMin , (m ,, pr))) } | inl (inl min<k) with TotalOrder.totality keyOrder min k
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = (yes (newMin , (m ,, pr))) } | inl (inl min<k) | inl (inl _) = lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} k v m
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = (yes (newMin , (m ,, pr))) } | inl (inl min<k) | inl (inr k<min) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) (PartialOrder.transitive (TotalOrder.order keyOrder) min<k k<min))
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = (yes (newMin , (m ,, pr))) } | inl (inl min<k) | inr min=k rewrite min=k = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) min<k)
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v m | inl (inr k<min) with TotalOrder.totality keyOrder k k
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v m | inl (inr k<min) | inl (inl k<k) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) k<k)
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v m | inl (inr k<min) | inl (inr k<k) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) k<k)
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v m | inl (inr k<min) | inr refl = refl
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v m | inr min=k with TotalOrder.totality keyOrder k k
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v m | inr min=k | inl (inl k<k) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) k<k)
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v m | inr min=k | inl (inr k<k) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) k<k)
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v m | inr min=k | inr refl with TotalOrder.totality keyOrder min k
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v m | inr min=k | inr refl | inl (inl min<k) rewrite min=k = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) min<k)
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v m | inr min=k | inr refl | inl (inr k<min) rewrite min=k = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) k<min)
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v m | inr refl | inr refl | inr refl = refl
lookupReducedSucceedsAfterUnrelatedAdd : {a b c : _} {keys : Set a} {values : Set b} {keyOrder : TotalOrder {_} {c} keys} {min : keys} (unrelatedK : keys) (unrelatedV : values) (k : keys) (v : values) ((TotalOrder._<_ keyOrder unrelatedK k) || (TotalOrder._<_ keyOrder k unrelatedK)) (m : ReducedMap keys values keyOrder min) (lookupReduced m k yes v) lookupReduced (addReducedMap unrelatedK unrelatedV m) k yes v
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v pr m lookupReducedSucceeds with TotalOrder.totality keyOrder min k'
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v pr m lookupReducedSucceeds | inl (inl min<k') with TotalOrder.totality keyOrder min k
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v pr record { firstEntry = firstEntry ; next = no } () | inl (inl min<k') | inl (inl min<k)
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v pr record { firstEntry = firstEntry ; next = (yes (a , (fst ,, snd))) } lookupReducedSucceeds | inl (inl min<k') | inl (inl min<k) = lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {a} k' v' k v pr fst lookupReducedSucceeds
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v (inl x) m lookupReducedSucceeds | inl (inl min<k') | inl (inr k<min) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) (PartialOrder.transitive (TotalOrder.order keyOrder) k<min (PartialOrder.transitive (TotalOrder.order keyOrder) min<k' x)))
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v (inr k<k') m () | inl (inl min<k') | inl (inr k<min)
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {.min} k' v' min v (inl x) m lookupReducedSucceeds | inl (inl min<k') | inr refl = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) (PartialOrder.transitive (TotalOrder.order keyOrder) x min<k'))
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {.min} k' v' min v (inr x) record { firstEntry = .v ; next = no } refl | inl (inl min<k') | inr refl = applyEquality yes refl
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {.min} k' v' min v (inr x) record { firstEntry = .v ; next = (yes (a , b)) } refl | inl (inl min<k') | inr refl = applyEquality yes refl
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v pr m lookupReducedSucceeds | inl (inr k'<min) with TotalOrder.totality keyOrder k' k
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v pr m lookupReducedSucceeds | inl (inr k'<min) | inl (inl k'<k) with TotalOrder.totality keyOrder min k
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v pr m lookupReducedSucceeds | inl (inr k'<min) | inl (inl k'<k) | inl (inl min<k) = lookupReducedSucceeds
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v pr m () | inl (inr k'<min) | inl (inl k'<k) | inl (inr k<min)
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v pr m lookupReducedSucceeds | inl (inr k'<min) | inl (inl k'<k) | inr refl = lookupReducedSucceeds
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v (inl x) m lookupReducedSucceeds | inl (inr k'<min) | inl (inr k<k') = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) (PartialOrder.transitive (TotalOrder.order keyOrder) x k<k'))
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v (inr x) m lookupReducedSucceeds | inl (inr k'<min) | inl (inr k<k') with TotalOrder.totality keyOrder min k
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v (inr _) m lookupReducedSucceeds | inl (inr k'<min) | inl (inr k<k') | inl (inl x) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) (PartialOrder.transitive (TotalOrder.order keyOrder) x (PartialOrder.transitive (TotalOrder.order keyOrder) k<k' k'<min)))
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v (inr _) m () | inl (inr k'<min) | inl (inr k<k') | inl (inr x)
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v (inr _) m lookupReducedSucceeds | inl (inr k'<min) | inl (inr k<k') | inr x rewrite x = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) (PartialOrder.transitive (TotalOrder.order keyOrder) k<k' k'<min))
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v (inl x) m lookupReducedSucceeds | inl (inr k'<min) | inr k'=k rewrite k'=k = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) x)
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v (inr x) m lookupReducedSucceeds | inl (inr k'<min) | inr k'=k rewrite k'=k = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) x)
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v pr m lookupReducedSucceeds | inr refl with TotalOrder.totality keyOrder min k
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {k'} k' v' k v (inl _) record { firstEntry = firstEntry ; next = no } () | inr refl | inl (inl min<k)
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {k'} k' v' k v (inl _) record { firstEntry = firstEntry ; next = (yes (a , (fst ,, snd))) } lookupReducedSucceeds | inr refl | inl (inl min<k) = lookupReducedSucceeds
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {k'} k' v' k v (inr x) m lookupReducedSucceeds | inr refl | inl (inl min<k) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) (PartialOrder.transitive (TotalOrder.order keyOrder) min<k x))
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {k'} k' v' k v (inl x) m () | inr refl | inl (inr k<min)
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {k'} k' v' k v (inr x) m () | inr refl | inl (inr k<min)
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {.min} .min v' min v (inl x) m lookupReducedSucceeds | inr refl | inr refl = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) x)
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {.min} .min v' min v (inr x) m lookupReducedSucceeds | inr refl | inr refl = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) x)
lookupReducedFailsAfterUnrelatedAdd : {a b c : _} {keys : Set a} {values : Set b} {keyOrder : TotalOrder {_} {c} keys} {min : keys} (unrelatedK : keys) (unrelatedV : values) (k : keys) ((TotalOrder._<_ keyOrder unrelatedK k) || (TotalOrder._<_ keyOrder k unrelatedK)) (m : ReducedMap keys values keyOrder min) (lookupReduced m k no) lookupReduced (addReducedMap unrelatedK unrelatedV m) k no
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k pr m lookupReducedFails with TotalOrder.totality keyOrder min k'
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k pr m lookupReducedFails | inl (inl min<k') with TotalOrder.totality keyOrder min k
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k pr record { firstEntry = firstEntry ; next = no } lookupReducedFails | inl (inl min<k') | inl (inl min<k) with TotalOrder.totality keyOrder k' k
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k pr record { firstEntry = firstEntry ; next = no } lookupReducedFails | inl (inl min<k') | inl (inl min<k) | inl (inl x) = refl
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k pr record { firstEntry = firstEntry ; next = no } lookupReducedFails | inl (inl min<k') | inl (inl min<k) | inl (inr x) = refl
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} .k v' k (inl x) record { firstEntry = firstEntry ; next = no } lookupReducedFails | inl (inl min<k') | inl (inl min<k) | inr refl = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) x)
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} .k v' k (inr x) record { firstEntry = firstEntry ; next = no } lookupReducedFails | inl (inl min<k') | inl (inl min<k) | inr refl = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) x)
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k pr record { firstEntry = firstEntry ; next = (yes (a , (fst ,, snd))) } lookupReducedFails | inl (inl min<k') | inl (inl min<k) = lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} k' v' k pr fst lookupReducedFails
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k pr m lookupReducedFails | inl (inl min<k') | inl (inr k<min) = refl
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {.min} k' v' min pr m () | inl (inl min<k') | inr refl
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k pr m lookupReducedFails | inl (inr k'<min) with TotalOrder.totality keyOrder min k
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k pr record { firstEntry = firstEntry ; next = next } lookupReducedFails | inl (inr k'<min) | inl (inl min<k) with TotalOrder.totality keyOrder k' k
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k pr record { firstEntry = firstEntry ; next = next } lookupReducedFails | inl (inr k'<min) | inl (inl min<k) | inl (inl k'<k) with TotalOrder.totality keyOrder min k
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k (inl x₁) record { firstEntry = firstEntry ; next = no } lookupReducedFails | inl (inr k'<min) | inl (inl min<k) | inl (inl k'<k) | inl (inl _) = refl
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k (inl x₁) record { firstEntry = firstEntry ; next = (yes (a , b)) } lookupReducedFails | inl (inr k'<min) | inl (inl min<k) | inl (inl k'<k) | inl (inl _) = lookupReducedFails
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k (inr x₁) record { firstEntry = firstEntry ; next = no } lookupReducedFails | inl (inr k'<min) | inl (inl min<k) | inl (inl k'<k) | inl (inl _) = refl
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k (inr x₁) record { firstEntry = firstEntry ; next = (yes (a , b)) } lookupReducedFails | inl (inr k'<min) | inl (inl min<k) | inl (inl k'<k) | inl (inl _) = lookupReducedFails
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k pr record { firstEntry = firstEntry ; next = next } lookupReducedFails | inl (inr k'<min) | inl (inl min<k) | inl (inl k'<k) | inl (inr x) = refl
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k pr record { firstEntry = firstEntry ; next = next } lookupReducedFails | inl (inr k'<min) | inl (inl min<k) | inl (inl k'<k) | inr refl = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) min<k)
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k pr record { firstEntry = firstEntry ; next = next } lookupReducedFails | inl (inr k'<min) | inl (inl min<k) | inl (inr k<k') = refl
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} .k v' k (inl x) record { firstEntry = firstEntry ; next = next } lookupReducedFails | inl (inr k'<min) | inl (inl min<k) | inr refl = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) x)
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} .k v' k (inr x) record { firstEntry = firstEntry ; next = next } lookupReducedFails | inl (inr k'<min) | inl (inl min<k) | inr refl = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) x)
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k pr m lookupReducedFails | inl (inr k'<min) | inl (inr k<min) with TotalOrder.totality keyOrder k' k
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k (inl _) m lookupReducedFails | inl (inr k'<min) | inl (inr k<min) | inl (inl k'<k) with TotalOrder.totality keyOrder min k
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k (inl _) m lookupReducedFails | inl (inr k'<min) | inl (inr k<min) | inl (inl k'<k) | inl (inl min<k) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) (PartialOrder.transitive (TotalOrder.order keyOrder) k<min min<k))
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k (inl _) m lookupReducedFails | inl (inr k'<min) | inl (inr k<min) | inl (inl k'<k) | inl (inr k<min') = refl
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k (inl _) m lookupReducedFails | inl (inr k'<min) | inl (inr k<min) | inl (inl k'<k) | inr refl = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) k<min)
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k (inr x) m lookupReducedFails | inl (inr k'<min) | inl (inr k<min) | inl (inl k'<k) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) (PartialOrder.transitive (TotalOrder.order keyOrder) k'<k x))
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k pr m lookupReducedFails | inl (inr k'<min) | inl (inr k<min) | inl (inr k<k') = refl
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} .k v' k (inl x) m lookupReducedFails | inl (inr k'<min) | inl (inr k<min) | inr refl = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) x)
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} .k v' k (inr x) m lookupReducedFails | inl (inr k'<min) | inl (inr k<min) | inr refl = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) x)
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k pr m () | inl (inr k'<min) | inr min=k
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} .min v' k pr m lookupReducedFails | inr refl with TotalOrder.totality keyOrder min k
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} .min v' k pr record { firstEntry = firstEntry ; next = no } lookupReducedFails | inr refl | inl (inl min<k) = refl
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} .min v' k pr record { firstEntry = firstEntry ; next = (yes (a , (fst ,, snd))) } lookupReducedFails | inr refl | inl (inl min<k) = lookupReducedFails
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} .min v' k pr m lookupReducedFails | inr refl | inl (inr k<min) = refl
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} .min v' k pr m () | inr refl | inr min=k
countReducedBehavesWhenAddingNotPresent : {a b c : _} {keys : Set a} {values : Set b} {keyOrder : TotalOrder {_} {c} keys} {min : keys} (k : keys) (v : values) (m : ReducedMap keys values keyOrder min) (lookupReduced {keyOrder = keyOrder} m k no) countReduced (addReducedMap k v m) succ (countReduced m)
countReducedBehavesWhenAddingNotPresent {keyOrder = keyOrder} {min} k v m lookupReducedFails with TotalOrder.totality keyOrder k min
countReducedBehavesWhenAddingNotPresent {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = next } lookupReducedFails | inl (inl k<min) with TotalOrder.totality keyOrder min k
countReducedBehavesWhenAddingNotPresent {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = next } lookupReducedFails | inl (inl k<min) | inl (inl min<k) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) (PartialOrder.transitive (TotalOrder.order keyOrder) min<k k<min))
countReducedBehavesWhenAddingNotPresent {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = next } lookupReducedFails | inl (inl k<min) | inl (inr _) = refl
countReducedBehavesWhenAddingNotPresent {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = next } lookupReducedFails | inl (inl k<min) | inr refl = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) k<min)
countReducedBehavesWhenAddingNotPresent {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = next } lookupReducedFails | inl (inr min<k) with TotalOrder.totality keyOrder min k
countReducedBehavesWhenAddingNotPresent {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = no } lookupReducedFails | inl (inr min<k) | inl (inl _) = refl
countReducedBehavesWhenAddingNotPresent {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = (yes (a , b)) } lookupReducedFails | inl (inr min<k) | inl (inl _) = applyEquality succ (countReducedBehavesWhenAddingNotPresent {keyOrder = keyOrder} k v (_&&_.fst b) lookupReducedFails)
countReducedBehavesWhenAddingNotPresent {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = next } lookupReducedFails | inl (inr min<k) | inl (inr k<min) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) (PartialOrder.transitive (TotalOrder.order keyOrder) min<k k<min))
countReducedBehavesWhenAddingNotPresent {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = next } lookupReducedFails | inl (inr min<k) | inr refl = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) min<k)
countReducedBehavesWhenAddingNotPresent {keyOrder = keyOrder} {min} k v m lookupReducedFails | inr refl with TotalOrder.totality keyOrder min min
countReducedBehavesWhenAddingNotPresent {keyOrder = keyOrder} {k} k v m lookupReducedFails | inr refl | inl (inl min<min) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) min<min)
countReducedBehavesWhenAddingNotPresent {keyOrder = keyOrder} {k} k v m lookupReducedFails | inr refl | inl (inr min<min) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) min<min)
countReducedBehavesWhenAddingNotPresent {keyOrder = keyOrder} {k} k v record { firstEntry = firstEntry ; next = no } () | inr refl | inr refl
countReducedBehavesWhenAddingNotPresent {keyOrder = keyOrder} {k} k v record { firstEntry = firstEntry ; next = (yes x) } () | inr refl | inr refl
countReducedBehavesWhenAddingPresent : {a b c : _} {keys : Set a} {values : Set b} {keyOrder : TotalOrder {_} {c} keys} {min : keys} (k : keys) (v v' : values) (m : ReducedMap keys values keyOrder min) (lookupReduced {keyOrder = keyOrder} m k yes v') countReduced (addReducedMap k v m) countReduced m
countReducedBehavesWhenAddingPresent {keyOrder = keyOrder} {min} k v v' m lookupReducedSucceeds with TotalOrder.totality keyOrder k min
countReducedBehavesWhenAddingPresent {keyOrder = keyOrder} {min} k v v' m lookupReducedSucceeds | inl (inl k<min) with TotalOrder.totality keyOrder min k
countReducedBehavesWhenAddingPresent {keyOrder = keyOrder} {min} k v v' m lookupReducedSucceeds | inl (inl k<min) | inl (inl min<k) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) (PartialOrder.transitive (TotalOrder.order keyOrder) k<min min<k))
countReducedBehavesWhenAddingPresent {keyOrder = keyOrder} {min} k v v' m () | inl (inl k<min) | inl (inr _)
countReducedBehavesWhenAddingPresent {keyOrder = keyOrder} {min} k v v' m lookupReducedSucceeds | inl (inl k<min) | inr refl = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) k<min)
countReducedBehavesWhenAddingPresent {keyOrder = keyOrder} {min} k v v' m lookupReducedSucceeds | inl (inr min<k) with TotalOrder.totality keyOrder min k
countReducedBehavesWhenAddingPresent {keyOrder = keyOrder} {min} k v v' record { firstEntry = firstEntry ; next = no } () | inl (inr min<k) | inl (inl _)
countReducedBehavesWhenAddingPresent {keyOrder = keyOrder} {min} k v v' record { firstEntry = firstEntry ; next = (yes (a , (fst ,, snd))) } lookupReducedSucceeds | inl (inr min<k) | inl (inl _) = applyEquality succ (countReducedBehavesWhenAddingPresent {keyOrder = keyOrder} k v v' fst lookupReducedSucceeds)
countReducedBehavesWhenAddingPresent {keyOrder = keyOrder} {min} k v v' m () | inl (inr min<k) | inl (inr k<min)
countReducedBehavesWhenAddingPresent {keyOrder = keyOrder} {min} k v v' m lookupReducedSucceeds | inl (inr min<k) | inr refl = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) min<k)
countReducedBehavesWhenAddingPresent {keyOrder = keyOrder} {min} k v v' m lookupReducedSucceeds | inr refl with TotalOrder.totality keyOrder min min
countReducedBehavesWhenAddingPresent {keyOrder = keyOrder} {k} k v v' m lookupReducedSucceeds | inr refl | inl (inl x) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) x)
countReducedBehavesWhenAddingPresent {keyOrder = keyOrder} {k} k v v' m lookupReducedSucceeds | inr refl | inl (inr x) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) x)
countReducedBehavesWhenAddingPresent {keyOrder = keyOrder} {k} k v v' record { firstEntry = firstEntry ; next = no } lookupReducedSucceeds | inr refl | inr refl = refl
countReducedBehavesWhenAddingPresent {keyOrder = keyOrder} {k} k v v' record { firstEntry = firstEntry ; next = (yes (a , b)) } lookupReducedSucceeds | inr refl | inr refl = refl
data Map {a b c : _} (keys : Set a) (values : Set b) (keyOrder : TotalOrder {_} {c} keys) : Set (a b c) where
empty : Map keys values keyOrder
nonempty : {min : keys} ReducedMap keys values keyOrder min Map keys values keyOrder
addMap : {a b c : _} {keys : Set a} {values : Set b} {keyOrder : TotalOrder {_} {c} keys} (m : Map keys values keyOrder) (k : keys) (v : values) Map keys values keyOrder
addMap empty k v = nonempty {min = k} record { firstEntry = v ; next = no }
addMap {keys = keys} {values} {keyOrder} (nonempty x) k v = nonempty (addReducedMap {keys = keys} {values} {keyOrder} k v x)
lookup : {a b c : _} {keys : Set a} {values : Set b} {keyOrder : TotalOrder {_} {c} keys} (m : Map keys values keyOrder) (target : keys) Maybe values
lookup {keyOrder = keyOrder} empty t = no
lookup {keyOrder = keyOrder} (nonempty x) t = lookupReduced x t
count : {a b c : _} {keys : Set a} {values : Set b} {keyOrder : TotalOrder {_} {c} keys} (m : Map keys values keyOrder)
count {keyOrder = keyOrder} empty = 0
count {keyOrder = keyOrder} (nonempty x) = countReduced x
keysReduced : {a b c : _} {keyDom : Set a} {values : Set b} {keyOrder : TotalOrder {_} {c} keyDom} {min : keyDom} (m : ReducedMap keyDom values keyOrder min) Vec keyDom (countReduced m)
keysReduced {min = min} record { firstEntry = firstEntry ; next = no } = min ,- []
keysReduced {min = min} record { firstEntry = firstEntry ; next = (yes (a , (fst ,, snd))) } = min ,- (keysReduced fst)
keys : {a b c : _} {keyDom : Set a} {values : Set b} {keyOrder : TotalOrder {_} {c} keyDom} (m : Map keyDom values keyOrder) Vec keyDom (count m)
keys empty = []
keys (nonempty m) = keysReduced m
lookupReducedWhenLess : {a b c : _} {keyDom : Set a} {values : Set b} {keyOrder : TotalOrder {_} {c} keyDom} {min : keyDom} (m : ReducedMap keyDom values keyOrder min) (k : keyDom) (TotalOrder._<_ keyOrder k min) (lookupReduced m k no)
lookupReducedWhenLess {keyOrder = keyOrder} {min} m k k<min with TotalOrder.totality keyOrder min k
lookupReducedWhenLess {keyOrder = keyOrder} {min} m k k<min | inl (inl min<k) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) (PartialOrder.transitive (TotalOrder.order keyOrder) k<min min<k))
lookupReducedWhenLess {keyOrder = keyOrder} {min} m k k<min | inl (inr _) = refl
lookupReducedWhenLess {keyOrder = keyOrder} {min} m k k<min | inr min=k rewrite min=k = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) k<min)
lookupCertainReduced : {a b c : _} {keyDom : Set a} {values : Set b} {keyOrder : TotalOrder {_} {c} keyDom} {min : keyDom} (m : ReducedMap keyDom values keyOrder min) (k : keyDom) (vecContains (keysReduced m) k) Sg values (λ v lookupReduced m k yes v)
lookupCertainReduced {keyDom = keyDom} {keyOrder = keyOrder} {min} record { firstEntry = firstEntry ; next = no } k pr = firstEntry , q
where
t : min k
t = vecSolelyContains k pr
q : lookupReduced {keys = keyDom} {keyOrder = keyOrder} {min} (record { firstEntry = firstEntry ; next = no }) k yes firstEntry
q with TotalOrder.totality keyOrder k k
q | inl (inl k<k) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) k<k)
q | inl (inr k<k) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) k<k)
q | inr refl with TotalOrder.totality keyOrder min k
q | inr refl | inl (inl min<k) rewrite t = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) min<k)
q | inr refl | inl (inr k<min) rewrite t = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) k<min)
q | inr refl | inr x = refl
lookupCertainReduced {keyOrder = keyOrder} {min} record { firstEntry = firstEntry ; next = (yes (a , (fst ,, snd))) } k pr with TotalOrder.totality keyOrder min k
lookupCertainReduced {keyOrder = keyOrder} {min} record { firstEntry = firstEntry ; next = (yes (a , (fst ,, snd))) } k record { index = zero ; index<m = _ ; isHere = isHere } | inl (inl min<k) rewrite isHere = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) min<k)
lookupCertainReduced {keyOrder = keyOrder} {min} record { firstEntry = firstEntry ; next = (yes (a , (fst ,, snd))) } k record { index = (succ index) ; index<m = index<m ; isHere = isHere } | inl (inl min<k) = lookupCertainReduced {keyOrder = keyOrder} fst k record { index = index ; index<m = canRemoveSuccFrom<N index<m ; isHere = isHere }
lookupCertainReduced {keyOrder = keyOrder} {min} record { firstEntry = firstEntry ; next = (yes (a , (fst ,, snd))) } k record { index = zero ; index<m = _ ; isHere = isHere } | inl (inr k<min) rewrite isHere = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) k<min)
lookupCertainReduced {keyOrder = keyOrder} {min} record { firstEntry = firstEntry ; next = (yes (a , (fst ,, snd))) } k record { index = (succ index) ; index<m = index<m ; isHere = isHere } | inl (inr k<min) with TotalOrder.totality keyOrder a k
lookupCertainReduced {keyOrder = keyOrder} {min} record { firstEntry = firstEntry ; next = (yes (a , (fst ,, snd))) } k record { index = (succ index) ; index<m = index<m ; isHere = isHere } | inl (inr k<min) | inl (inl a<k) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) (PartialOrder.transitive (TotalOrder.order keyOrder) k<min (PartialOrder.transitive (TotalOrder.order keyOrder) snd a<k)))
lookupCertainReduced {values = values} {keyOrder = keyOrder} {min} record { firstEntry = firstEntry ; next = (yes (a , (fst ,, snd))) } k record { index = (succ index) ; index<m = index<m ; isHere = isHere } | inl (inr k<min) | inl (inr k<a) = exFalso h
where
f : Sg values (λ v lookupReduced fst k yes v)
f = lookupCertainReduced {keyOrder = keyOrder} fst k record { index = index ; index<m = canRemoveSuccFrom<N index<m ; isHere = isHere }
g : lookupReduced fst k no
g = lookupReducedWhenLess fst k k<a
noIsNotYes : {a : _} {A : Set a} {b : A} (no yes b) False
noIsNotYes {a} {A} {b} ()
h : False
h with f
h | a , b rewrite g = noIsNotYes b
lookupCertainReduced {keyOrder = keyOrder} {min} record { firstEntry = firstEntry ; next = (yes (a , (fst ,, snd))) } k record { index = (succ index) ; index<m = index<m ; isHere = isHere } | inl (inr k<min) | inr refl = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) (PartialOrder.transitive (TotalOrder.order keyOrder) snd k<min))
lookupCertainReduced {keyOrder = keyOrder} {min} record { firstEntry = firstEntry ; next = (yes (a , (fst ,, snd))) } k pr | inr min=k = firstEntry , refl
lookupCertain : {a b c : _} {keyDom : Set a} {values : Set b} {keyOrder : TotalOrder {_} {c} keyDom} (m : Map keyDom values keyOrder) (k : keyDom) (vecContains (keys m) k) Sg values (λ v lookup m k yes v)
lookupCertain {keyOrder = keyOrder} empty k record { index = index ; index<m = (le x ()) ; isHere = isHere }
lookupCertain {keyOrder = keyOrder} (nonempty {min} m) k pr = lookupCertainReduced {keyOrder = keyOrder} {min} m k pr

20
KeyValueWithDomain.agda Normal file
View File

@@ -0,0 +1,20 @@
{-# OPTIONS --warning=error --safe #-}
open import LogicalFormulae
open import Orders
open import Maybe
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import KeyValue
open import Vectors
open import Naturals
module KeyValueWithDomain where
record MapWithDomain {a b c : _} (keyDom : Set a) (values : Set b) (keyOrder : TotalOrder {_} {c} keyDom) : Set (a b c) where
field
map : Map keyDom values keyOrder
domain : Vec keyDom (count map)
domainIsIt : keys map domain
lookup' : (key : keyDom) (vecContains domain key) Sg values (λ v lookup map key yes v)
lookup' k cont rewrite equalityCommutative domainIsIt = lookupCertain {keyOrder = keyOrder} map k cont

58
Lists.agda Normal file
View File

@@ -0,0 +1,58 @@
{-# OPTIONS --safe --warning=error #-}
open import LogicalFormulae
open import Naturals -- for length
module Lists where
data Vec {a} (A : Set a) : Set a where
[] : Vec A 0
_::_ : {len : } (x : A) (xs : Vec A len) Vec A (succ len)
[_] : {a : _} {A : Set a} (a : A) Vec A 1
[ a ] = a :: []
_++_ : {a : _} {A : Set a} {m n : } Vec A m Vec A n Vec A (m +N n)
[] ++ m = m
(x :: l) ++ m = x :: (l ++ m)
appendEmptyList : {a : _} {A : Set a} {m : } (l : Vec A m) (l ++ [] l)
appendEmptyList [] = refl
appendEmptyList (x :: l) = applyEquality (_::_ x) (appendEmptyList l)
concatAssoc : {a : _} {A : Set a} {m n o : } (x : Vec A m) (y : Vec A n) (z : Vec A o) ((x ++ y) ++ z) x ++ (y ++ z)
concatAssoc [] m n = refl
concatAssoc (x :: l) m n = applyEquality (_::_ x) (concatAssoc l m n)
canMovePrepend : {a : _} {A : Set a} (l : A) {m n : } (x : Vec A m) (y : Vec A n) ((l :: x) ++ y l :: (x ++ y))
canMovePrepend l [] n = refl
canMovePrepend l (x :: m) n = refl
rev : {a : _} {A : Set a} {m : } Vec A m Vec A m
rev [] = []
rev (x :: l) = (rev l) ++ [ x ]
revIsHom : {a : _} {A : Set a} {m n : } (l1 : Vec A m) (l2 : Vec A n) (rev (l1 ++ l2) (rev l2) ++ (rev l1))
revIsHom l1 [] = applyEquality rev (appendEmptyList l1)
revIsHom [] (x :: l2) with (rev l2 ++ [ x ])
... | r = equalityCommutative (appendEmptyList r)
revIsHom (w :: l1) (x :: l2) = transitivity t (equalityCommutative s)
where
s : ((rev l2 ++ [ x ]) ++ (rev l1 ++ [ w ])) (((rev l2 ++ [ x ]) ++ rev l1) ++ [ w ])
s = equalityCommutative (concatAssoc (rev l2 ++ (x :: [])) (rev l1) ([ w ]))
t' : rev (l1 ++ (x :: l2)) rev (x :: l2) ++ rev l1
t' = revIsHom l1 (x :: l2)
t : (rev (l1 ++ (x :: l2)) ++ [ w ]) ((rev l2 ++ [ x ]) ++ rev l1) ++ [ w ]
t = applyEquality (λ r r ++ [ w ]) {rev (l1 ++ (x :: l2))} {((rev l2) ++ [ x ]) ++ rev l1} (transitivity t' (applyEquality (λ r r ++ rev l1) {rev l2 ++ (x :: [])} {rev l2 ++ (x :: [])} refl))
revRevIsId : {a : _} {A : Set a} {m : } (l : Vec A m) (rev (rev l) l)
revRevIsId [] = refl
revRevIsId (x :: l) = t
where
s : rev (rev l ++ [ x ] ) [ x ] ++ rev (rev l)
s = revIsHom (rev l) [ x ]
t : rev (rev l ++ [ x ] ) [ x ] ++ l
t = identityOfIndiscernablesRight (rev (rev l ++ (x :: []))) ([ x ] ++ rev (rev l)) ([ x ] ++ l) _≡_ s (applyEquality (λ n [ x ] ++ n) (revRevIsId l))
map : {a : _} {b : _} {A : Set a} {B : Set b} {m : } (f : A B) Vec A m Vec B m
map f [] = []
map f (x :: list) = (f x) :: (map f list)

112
LogicalFormulae.agda Normal file
View File

@@ -0,0 +1,112 @@
{-# OPTIONS --safe --warning=error #-}
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
module LogicalFormulae where
infix 5 _≡_
data _≡_ {a} {A : Set a} (x : A) : A Set a where
refl : x x
{-# BUILTIN EQUALITY _≡_ #-}
data False : Set where
record True : Set where
record Unit : Set where
infix 10 _||_
data _||_ {a b} (A : Set a) (B : Set b) : Set (a b) where
inl : A A || B
inr : B A || B
data Bool : Set where
BoolTrue : Bool
BoolFalse : Bool
infix 15 _&&_
record _&&_ {a b} (A : Set a) (B : Set b) : Set (a b) where
constructor _,,_
field
fst : A
snd : B
infix 15 _&_&_
record _&_&_ {a b c} (A : Set a) (B : Set b) (C : Set c) : Set (a b c) where
field
one : A
two : B
three : C
data Sg {n : _} {m : _} (A : Set m) (B : A Set n) : Set (m n) where
_,_ : (a : A) (b : B a) Sg A B
underlying : {n m : _} {A : _} {prop : _} Sg {n} {m} A prop -> A
underlying (a , b) = a
transitivity : {a : _} {A : Set a} {x y z : A} (x y) (y z) (x z)
transitivity refl refl = refl
liftEquality : {a : _} {A B : Set a} (f g : A B) (x y : A) (f g) (x y) ((f x) (g y))
liftEquality f .f x .x refl refl = refl
applyEquality : {a : _} {b : _} {A : Set a} {B : Set b} (f : A B) {x y : A} (x y) ((f x) (f y))
applyEquality {A} {B} f {x} {.x} refl = refl
identityOfIndiscernablesLeft : {m n o : _} {A : Set m} {B : Set n} (a : A) (b : B) (c : A) (prop : A B Set o) (prop a b) (a c) (prop c b)
identityOfIndiscernablesLeft a b .a prop pAB refl = pAB
identityOfIndiscernablesRight : {m n o : _} {A : Set m} {B : Set n} (a : A) (b c : B) (prop : A B Set o) (prop a b) (b c) (prop a c)
identityOfIndiscernablesRight a b .b prop prAB refl = prAB
equalityCommutative : {a : _} {A : Set a} {x y : A} (x y) (y x)
equalityCommutative refl = refl
exFalso : {n : _} {A : Set n} False A
exFalso {a} = λ ()
orIsAssociative : {n : _} {a b c : Set n} ((a || b) || c) (a || (b || c))
orIsAssociative (inl (inl x)) = inl x
orIsAssociative (inl (inr x)) = inr (inl x)
orIsAssociative (inr x) = inr (inr x)
leqConstructive : {n : _} {p : Set n} (p || (p False)) (((p False) False) p)
leqConstructive (inl p) = λ _ p
leqConstructive (inr notP) = λ negnegP exFalso (negnegP notP)
lemConverse : {n : _} {p : Set n} p ((p False) False)
lemConverse p = λ notP notP p
if_then_else_ : {a : _} {A : Set a} Bool A A A
if BoolTrue then tr else fls = tr
if BoolFalse then tr else fls = fls
not : Bool Bool
not BoolTrue = BoolFalse
not BoolFalse = BoolTrue
typeCast : {a : _} {A : Set a} {B : Set a} (el : A) (pr : A B) B
typeCast {a} {A} {.A} elt refl = elt
typeCastEqual : {a : _} {A : Set a} {B : Set a} {el : A} (pr1 pr2 : A B) (typeCast el pr1) (typeCast el pr2)
typeCastEqual {a} {A} {.A} {el} refl refl = refl
data Singleton {a} {A : Set a} (x : A) : Set a where
_with≡_ : (y : A) x y Singleton x
inspect : {a} {A : Set a} (x : A) Singleton x
inspect x = x with refl
curry : {a b c : _} {A : Set a} {B : Set b} {C : Set c} (f : A B C) ((Sg A (λ _ B)) C)
curry f (a , b) = f a b
uncurry : {a b c : _} {A : Set a} {B : Set b} {C : Set c} (f : (Sg A (λ _ B)) C) (A B C)
uncurry f a b = f (a , b)
decidableOr : {a b : _} (A : Set a) (B : Set b) (A || (A False)) (A || B) (A || ((A False) && B))
decidableOr {a} {b} A B decidable (inl x) = inl x
decidableOr {a} {b} A B (inl y) (inr x) = inl y
decidableOr {a} {b} A B (inr y) (inr x) = inr (record { fst = y ; snd = x})
reflRefl : {a : _} {A : Set a} {r s : A} (pr1 pr2 : r s) pr1 pr2
reflRefl refl refl = refl

View File

@@ -0,0 +1,8 @@
{-# OPTIONS --safe --warning=error #-}
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import LogicalFormulae
module LogicalFormulaeProofs where
transitivity : LogicalFormulae.transitivity
transitivity refl refl = refl

33
Maybe.agda Normal file
View File

@@ -0,0 +1,33 @@
{-# OPTIONS --warning=error --safe #-}
open import LogicalFormulae
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
module Maybe where
data Maybe {a : _} (A : Set a) : Set a where
no : Maybe A
yes : A Maybe A
joinMaybe : {a : _} {A : Set a} Maybe (Maybe A) Maybe A
joinMaybe no = no
joinMaybe (yes s) = s
bindMaybe : {a b : _} {A : Set a} {B : Set b} Maybe A (A Maybe B) Maybe B
bindMaybe no f = no
bindMaybe (yes x) f = f x
applyMaybe : {a b : _} {A : Set a} {B : Set b} Maybe (A B) Maybe A Maybe B
applyMaybe f no = no
applyMaybe no (yes x) = no
applyMaybe (yes f) (yes x) = yes (f x)
yesInjective : {a : _} {A : Set a} {x y : A} (yes x yes y) x y
yesInjective {a} {A} {x} {.x} refl = refl
mapMaybe : {a b : _} {A : Set a} {B : Set b} (f : A B) Maybe A Maybe B
mapMaybe f no = no
mapMaybe f (yes x) = yes (f x)
defaultValue : {a : _} {A : Set a} (default : A) Maybe A A
defaultValue default no = default
defaultValue default (yes x) = x

842
Naturals.agda Normal file
View File

@@ -0,0 +1,842 @@
{-# OPTIONS --warning=error --safe #-}
open import LogicalFormulae
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import WellFoundedInduction
open import Functions
open import Orders
module Naturals where
data : Set where
zero :
succ :
{-# BUILTIN NATURAL #-}
transitivityN = transitivity {_} {}
infix 15 _+N_
infix 100 succ
_+N_ :
zero +N y = y
succ x +N y = succ (x +N y)
infix 25 _*N_
_*N_ :
zero *N y = zero
(succ x) *N y = y +N (x *N y)
data _even : Set where
zero_is_even : zero even
add_two_stays_even : x x even succ (succ x) even
four_is_even : succ (succ (succ (succ zero))) even
four_is_even = add_two_stays_even (succ (succ zero)) (add_two_stays_even zero zero_is_even)
infix 5 _<NLogical_
_<NLogical_ : Set
zero <NLogical zero = False
zero <NLogical (succ n) = True
(succ n) <NLogical zero = False
(succ n) <NLogical (succ m) = n <NLogical m
infix 5 _<N_
record _<N_ (a : ) (b : ) : Set where
constructor le
field
x :
proof : (succ x) +N a b
infix 5 _≤N_
_≤N_ : Set
a ≤N b = (a <N b) || (a b)
addzeroLeft : (x : ) (zero +N x) x
addzeroLeft x = refl
addZeroRight : (x : ) (x +N zero) x
addZeroRight zero = refl
addZeroRight (succ x) = applyEquality succ (addZeroRight x)
succExtracts : (x y : ) (x +N succ y) (succ (x +N y))
succExtracts zero y = refl
succExtracts (succ x) y = applyEquality succ (succExtracts x y)
succCanMove : (x y : ) (x +N succ y) (succ x +N y)
succCanMove x y = transitivity (succExtracts x y) refl
aLessSucc : (a : ) (a <NLogical succ a)
aLessSucc zero = record {}
aLessSucc (succ a) = aLessSucc a
succPreservesInequalityLogical : {a b : } (a <NLogical b) (succ a <NLogical succ b)
succPreservesInequalityLogical {a} {b} prAB = prAB
lessTransitiveLogical : {a b c : } (a <NLogical b) (b <NLogical c) (a <NLogical c)
lessTransitiveLogical {a} {zero} {zero} prAB prBC = prAB
lessTransitiveLogical {a} {(succ b)} {zero} prAB ()
lessTransitiveLogical {zero} {zero} {(succ c)} prAB prBC = record {}
lessTransitiveLogical {(succ a)} {zero} {(succ c)} () prBC
lessTransitiveLogical {zero} {(succ b)} {(succ c)} prAB prBC = record {}
lessTransitiveLogical {(succ a)} {(succ b)} {(succ c)} prAB prBC = lessTransitiveLogical {a} {b} {c} prAB prBC
aLessXPlusSuccA : (a x : ) (a <NLogical x +N succ a)
aLessXPlusSuccA a zero = aLessSucc a
aLessXPlusSuccA zero (succ x) = record {}
aLessXPlusSuccA (succ a) (succ x) = lessTransitiveLogical {a} {succ a} {x +N succ (succ a)} (aLessXPlusSuccA a zero) (aLessXPlusSuccA (succ a) x)
leImpliesLogical<N : {a b : } (a <N b) (a <NLogical b)
leImpliesLogical<N {zero} {zero} (le x ())
leImpliesLogical<N {zero} {(succ b)} (le x proof) = record {}
leImpliesLogical<N {(succ a)} {zero} (le x ())
leImpliesLogical<N {(succ a)} {(succ .(succ a))} (le zero refl) = aLessSucc a
leImpliesLogical<N {(succ a)} {(succ .(succ (x +N succ a)))} (le (succ x) refl) = succPreservesInequalityLogical {a} {succ (x +N succ a)} (lessTransitiveLogical {a} {succ a} {succ (x +N succ a)} (aLessSucc a) (aLessXPlusSuccA a x))
logical<NImpliesLe : {a b : } (a <NLogical b) (a <N b)
logical<NImpliesLe {zero} {zero} ()
_<N_.x (logical<NImpliesLe {zero} {succ b} prAB) = b
_<N_.proof (logical<NImpliesLe {zero} {succ b} prAB) = applyEquality succ (addZeroRight b)
logical<NImpliesLe {(succ a)} {zero} ()
logical<NImpliesLe {(succ a)} {(succ b)} prAB with logical<NImpliesLe {a} prAB
logical<NImpliesLe {(succ a)} {(succ .(succ (x +N a)))} prAB | le x refl = le x (succCanMove (succ x) a)
lessTransitive : {a b c : } (a <N b) (b <N c) (a <N c)
lessTransitive {a} {b} {c} prab prbc = logical<NImpliesLe (lessTransitiveLogical {a} {b} {c} (leImpliesLogical<N prab) (leImpliesLogical<N prbc))
canRemoveSuccFrom<N : {a b : } (succ a <N succ b) (a <N b)
canRemoveSuccFrom<N {a} {b} prAB = logical<NImpliesLe (leImpliesLogical<N prAB)
succPreservesInequality : {a b : } (a <N b) (succ a <N succ b)
succPreservesInequality {a} {b} prAB = logical<NImpliesLe (succPreservesInequalityLogical {a} {b} (leImpliesLogical<N prAB))
lessIrreflexive : {a : } (a <N a) False
lessIrreflexive {zero} pr = leImpliesLogical<N pr
lessIrreflexive {succ a} pr = lessIrreflexive {a} (logical<NImpliesLe {a} {a} (leImpliesLogical<N {succ a} {succ a} pr))
additionNIsCommutative : (x y : ) (x +N y) (y +N x)
additionNIsCommutative zero y = transitivity (addzeroLeft 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)))
succInjective : {a b : } (succ a succ b) a b
succInjective {a} {.a} refl = refl
addingPreservesEqualityRight : {a b : } (c : ) (a b) (a +N c b +N c)
addingPreservesEqualityRight {a} {b} c pr = applyEquality (λ n -> n +N c) pr
addingPreservesEqualityLeft : {a b : } (c : ) (a b) (c +N a c +N b)
addingPreservesEqualityLeft {a} {b} c pr = applyEquality (λ n -> c +N n) pr
additionNIsAssociative : (a b c : ) ((a +N b) +N c) (a +N (b +N c))
additionNIsAssociative zero b c = refl
additionNIsAssociative (succ a) zero c = transitivity (transitivity (applyEquality (λ n n +N c) (applyEquality succ (addZeroRight a))) refl) (transitivity refl refl)
additionNIsAssociative (succ a) (succ b) c = transitivity refl (transitivity refl (transitivity (applyEquality succ (additionNIsAssociative a (succ b) c)) refl))
productZeroIsZeroLeft : (a : ) (zero *N a zero)
productZeroIsZeroLeft a = refl
productZeroIsZeroRight : (a : ) (a *N zero zero)
productZeroIsZeroRight zero = refl
productZeroIsZeroRight (succ a) = productZeroIsZeroRight a
productWithOneLeft : (a : ) ((succ zero) *N a) a
productWithOneLeft a = transitivity refl (transitivity (applyEquality (λ { m -> a +N m }) refl) (additionNIsCommutative a zero))
productWithOneRight : (a : ) a *N succ zero a
productWithOneRight zero = refl
productWithOneRight (succ a) = transitivity refl (addingPreservesEqualityLeft (succ zero) (productWithOneRight a))
productDistributes : (a b c : ) (a *N (b +N c)) a *N b +N a *N c
productDistributes zero b c = refl
productDistributes (succ a) b c = transitivity refl
(transitivity (addingPreservesEqualityLeft (b +N c) (productDistributes a b c))
(transitivity (equalityCommutative (additionNIsAssociative (b +N c) (a *N b) (a *N c)))
(transitivity (addingPreservesEqualityRight (a *N c) (additionNIsCommutative (b +N c) (a *N b)))
(transitivity (addingPreservesEqualityRight (a *N c) (equalityCommutative (additionNIsAssociative (a *N b) b c)))
(transitivity (addingPreservesEqualityRight (a *N c) (addingPreservesEqualityRight c (additionNIsCommutative (a *N b) b)))
(transitivity (addingPreservesEqualityRight (a *N c) (addingPreservesEqualityRight {b +N a *N b} {(succ a) *N b} c (refl)))
(transitivity (additionNIsAssociative ((succ a) *N b) c (a *N c))
(transitivity (addingPreservesEqualityLeft (succ a *N b) refl)
refl)
)
))))))
succIsAddOne : (a : ) succ a a +N succ zero
succIsAddOne a = equalityCommutative (transitivity (additionNIsCommutative a (succ zero)) refl)
productPreservesEqualityLeft : (a : ) {b c : } b c a *N b a *N c
productPreservesEqualityLeft a {b} {.b} refl = refl
aSucB : (a b : ) a *N succ b a *N b +N a
aSucB a b = transitivity {_} {} {a *N succ b} {a *N (b +N succ zero)} (productPreservesEqualityLeft a (succIsAddOne b)) (transitivity {_} {} {a *N (b +N succ zero)} {a *N b +N a *N succ zero} (productDistributes a b (succ zero)) (addingPreservesEqualityLeft (a *N b) (productWithOneRight a)))
aSucBRight : (a b : ) (succ a) *N b a *N b +N b
aSucBRight a b = additionNIsCommutative b (a *N b)
multiplicationNIsCommutative : (a b : ) (a *N b) (b *N a)
multiplicationNIsCommutative zero b = transitivity (productZeroIsZeroLeft b) (equalityCommutative (productZeroIsZeroRight b))
multiplicationNIsCommutative (succ a) zero = multiplicationNIsCommutative a zero
multiplicationNIsCommutative (succ a) (succ b) = transitivityN refl
(transitivityN (addingPreservesEqualityLeft (succ b) (aSucB a b))
(transitivityN {succ b +N (a *N b +N a)} {(a *N b +N a) +N succ b} (additionNIsCommutative (succ b) (a *N b +N a))
(transitivityN {(a *N b +N a) +N succ b} {a *N b +N (a +N succ b)} (additionNIsAssociative (a *N b) a (succ b))
(transitivityN {a *N b +N (a +N succ b)} {a *N b +N ((succ a) +N b)} (addingPreservesEqualityLeft (a *N b) (succCanMove a b))
(transitivityN {a *N b +N ((succ a) +N b)} {a *N b +N (b +N succ a)} (addingPreservesEqualityLeft (a *N b) (additionNIsCommutative (succ a) b))
(transitivityN {a *N b +N (b +N succ a)} {(a *N b +N b) +N succ a} (equalityCommutative (additionNIsAssociative (a *N b) b (succ a)))
(transitivity (addingPreservesEqualityRight (succ a) (equalityCommutative (aSucBRight a b)))
(transitivity (addingPreservesEqualityRight (succ a) (multiplicationNIsCommutative (succ a) b))
(transitivity (additionNIsCommutative (b *N (succ a)) (succ a))
refl
)))))))))
productDistributes' : (a b c : ) (a +N b) *N c a *N c +N b *N c
productDistributes' a b c rewrite multiplicationNIsCommutative (a +N b) c | productDistributes c a b | multiplicationNIsCommutative c a | multiplicationNIsCommutative c b = refl
flipProductsWithinSum : (a b c : ) (c *N a +N c *N b a *N c +N b *N c)
flipProductsWithinSum a b c = transitivity (addingPreservesEqualityRight (c *N b) (multiplicationNIsCommutative c a)) ((addingPreservesEqualityLeft (a *N c) (multiplicationNIsCommutative c b)))
productDistributesRight : (a b c : ) (a +N b) *N c a *N c +N b *N c
productDistributesRight a b c = transitivity (multiplicationNIsCommutative (a +N b) c) (transitivity (productDistributes c a b) (flipProductsWithinSum a b c))
multiplicationNIsAssociative : (a b c : ) (a *N (b *N c)) ((a *N b) *N c)
multiplicationNIsAssociative zero b c = refl
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)))
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)
where
g : result2 0
g = canSubtractFromEqualityLeft {a} {_} {0} (transitivity pr (equalityCommutative (addZeroRight a)))
g' : result1 0
g' = canSubtractFromEqualityLeft {a} {_} {0} (transitivity pr1 (equalityCommutative (addZeroRight a)))
-N : {a : } {b : } (pr : a ≤N b) subtractionNResult a b pr
-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) | 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) }
addOneToWeakInequality : {a b : } (a ≤N b) (succ a ≤N succ b)
addOneToWeakInequality {a} {b} (inl ineq) = inl (succPreservesInequality ineq)
addOneToWeakInequality {a} {.a} (inr refl) = inr refl
bumpUpSubtraction : {a b : } (p1 : a ≤N b) (s : subtractionNResult a b p1) Sg (subtractionNResult (succ a) (succ b) (addOneToWeakInequality p1)) (λ n subtractionNResult.result n subtractionNResult.result s)
bumpUpSubtraction {a} {b} a<=b record { result = result ; pr = pr } = record { result = result ; pr = applyEquality succ pr } , refl
addMinus : {a : } {b : } (pr : a ≤N b) subtractionNResult.result (-N {a} {b} pr) +N a b
addMinus {zero} {zero} p = refl
addMinus {zero} {succ b} pr = applyEquality succ (addZeroRight b)
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
where
pr'' : (a <N b) || (a b)
pr'' = (inl (le (_<N_.x x) (transitivity (equalityCommutative (succExtracts (_<N_.x x) a)) (succInjective (_<N_.proof x)))))
previous : subtractionNResult a b pr''
previous = -N pr''
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)
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)))
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)))
productNonzeroIsNonzero : {a b : } zero <N a zero <N b zero <N a *N b
productNonzeroIsNonzero {zero} {b} prA prB = prA
productNonzeroIsNonzero {succ a} {zero} prA ()
productNonzeroIsNonzero {succ a} {succ b} prA prB = logical<NImpliesLe (record {})
multiplyIncreases : (a : ) (b : ) succ zero <N a zero <N b b <N a *N b
multiplyIncreases zero b (le x ()) prB
multiplyIncreases (succ zero) b prA prb with leImpliesLogical<N {succ zero} {succ zero} prA
multiplyIncreases (succ zero) b prA prb | ()
multiplyIncreases (succ (succ a)) zero prA ()
multiplyIncreases (succ (succ a)) (succ b) prA prB =
identityOfIndiscernablesRight (succ b) (succ b +N (succ a) *N succ b) (succ (succ a) *N succ b) (λ a b a <N b) k refl
where
h : zero <N (succ a) *N (succ b)
h = productNonzeroIsNonzero {succ a} {succ b} (logical<NImpliesLe (record {})) (logical<NImpliesLe (record {}))
i : zero +N succ b <N (succ a) *N (succ b) +N succ b
i = additionPreservesInequality {zero} {succ a *N succ b} (succ b) h
j : zero +N succ b <N succ b +N (succ a *N succ b)
j = identityOfIndiscernablesRight (zero +N succ b) (succ ((b +N a *N succ b) +N succ b)) (succ (b +N succ (b +N a *N succ b))) (λ a b a <N b) i (additionNIsCommutative (succ (b +N a *N succ b)) (succ b))
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
productCancelsRight (succ a) zero (succ c) aPos eq = contr
where
h : zero succ c *N succ a
h = eq
contr : zero succ c
contr = naughtE h
productCancelsRight zero (succ b) zero (le x ()) eq
productCancelsRight (succ a) (succ b) zero aPos eq = contr
where
h : succ b *N succ a zero
h = eq
contr : succ b zero
contr = naughtE (equalityCommutative h)
productCancelsRight zero (succ b) (succ c) (le x ()) eq
productCancelsRight (succ a) (succ b) (succ c) aPos eq = applyEquality succ (productCancelsRight (succ a) b c aPos l)
where
i : succ a +N b *N succ a succ c *N succ a
i = eq
j : succ c *N succ a succ a +N c *N succ a
j = refl
k : succ a +N b *N succ a succ a +N c *N succ a
k = transitivity i j
l : b *N succ a c *N succ a
l = canSubtractFromEqualityLeft {succ a} {b *N succ a} {c *N succ a} k
productCancelsLeft : (a b c : ) (zero <N a) (a *N b a *N c) (b c)
productCancelsLeft a b c aPos pr = productCancelsRight a b c aPos j
where
i : b *N a a *N c
i = identityOfIndiscernablesLeft (a *N b) (a *N c) (b *N a) _≡_ pr (multiplicationNIsCommutative a b)
j : b *N a c *N a
j = identityOfIndiscernablesRight (b *N a) (a *N c) (c *N a) _≡_ i (multiplicationNIsCommutative a c)
productCancelsRight' : (a b c : ) (b *N a c *N a) (a zero) || (b c)
productCancelsRight' zero b c pr = inl refl
productCancelsRight' (succ a) b c pr = inr (productCancelsRight (succ a) b c (succIsPositive a) pr)
productCancelsLeft' : (a b c : ) (a *N b a *N c) (a zero) || (b c)
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)))
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))
canTimesOneOnRight : {a b : } (a <N b) a <N (b *N (succ zero))
canTimesOneOnRight {a} {b} prAB = identityOfIndiscernablesRight a b (b *N (succ zero)) (λ x y x <N y) prAB (equalityCommutative (productWithOneRight b))
canSwapAddOnLeftOfInequality : {a b c : } (a +N b <N c) (b +N a <N c)
canSwapAddOnLeftOfInequality {a} {b} {c} pr = identityOfIndiscernablesLeft (a +N b) c (b +N a) (λ x y x <N y) pr (additionNIsCommutative a b)
canSwapAddOnRightOfInequality : {a b c : } (a <N b +N c) (a <N c +N b)
canSwapAddOnRightOfInequality {a} {b} {c} pr = identityOfIndiscernablesRight a (b +N c) (c +N b) (λ x y x <N y) pr (additionNIsCommutative b c)
naturalsAreNonnegative : (a : ) (a <NLogical zero) False
naturalsAreNonnegative zero pr = pr
naturalsAreNonnegative (succ x) pr = pr
canFlipMultiplicationsInIneq : {a b c d : } (a *N b <N c *N d) b *N a <N d *N c
canFlipMultiplicationsInIneq {a} {b} {c} {d} pr = identityOfIndiscernablesRight (b *N a) (c *N d) (d *N c) _<N_ (identityOfIndiscernablesLeft (a *N b) (c *N d) (b *N a) _<N_ pr (multiplicationNIsCommutative a b)) (multiplicationNIsCommutative c d)
bumpDownOnRight : (a c : ) c *N succ a c *N a +N c
bumpDownOnRight a c = transitivity (multiplicationNIsCommutative c (succ a)) (transitivity refl (transitivity (additionNIsCommutative c (a *N c)) ((addingPreservesEqualityRight c (multiplicationNIsCommutative a c) ))))
lessRespectsMultiplicationLeft : (a b c : ) (a <N b) (zero <N c) (c *N a <N c *N b)
lessRespectsMultiplicationLeft zero zero c (le x ()) cPos
lessRespectsMultiplicationLeft zero (succ b) zero prAB (le x ())
lessRespectsMultiplicationLeft zero (succ b) (succ c) prAB cPos = i
where
j : zero <N succ c *N succ b
j = productNonzeroIsNonzero cPos prAB
i : succ c *N zero <N succ c *N succ b
i = identityOfIndiscernablesLeft zero (succ c *N (succ b)) ((succ c) *N zero) _<N_ j (equalityCommutative (productZeroIsZeroRight c))
lessRespectsMultiplicationLeft (succ a) zero c (le x ()) cPos
lessRespectsMultiplicationLeft (succ a) (succ b) c prAB cPos = m
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
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
m = identityOfIndiscernablesRight (c *N succ a) (c *N b +N c) (c *N succ b) _<N_ i (equalityCommutative (bumpDownOnRight b c))
lessRespectsMultiplication : (a b c : ) (a <N b) (zero <N c) (a *N c <N b *N c)
lessRespectsMultiplication a b c prAB cPos = canFlipMultiplicationsInIneq {c} {a} {c} {b} (lessRespectsMultiplicationLeft a b c prAB cPos)
succIsNonzero : {a : } (succ a zero) False
succIsNonzero {a} ()
orderIsTotal : (a b : ) ((a <N b) || (b <N a)) || (a b)
orderIsTotal zero zero = inr refl
orderIsTotal zero (succ b) = inl (inl (logical<NImpliesLe (record {})))
orderIsTotal (succ a) zero = inl (inr (logical<NImpliesLe (record {})))
orderIsTotal (succ a) (succ b) with orderIsTotal a b
orderIsTotal (succ a) (succ b) | inl (inl x) = inl (inl (logical<NImpliesLe (leImpliesLogical<N x)))
orderIsTotal (succ a) (succ b) | inl (inr x) = inl (inr (logical<NImpliesLe (leImpliesLogical<N x)))
orderIsTotal (succ a) (succ b) | inr x = inr (applyEquality succ x)
orderIsIrreflexive : {a b : } (a <N b) (b <N a) False
orderIsIrreflexive {zero} {b} prAB (le x ())
orderIsIrreflexive {(succ a)} {zero} (le x ()) prBA
orderIsIrreflexive {(succ a)} {(succ b)} prAB prBA = orderIsIrreflexive {a} {b} (logical<NImpliesLe (leImpliesLogical<N prAB)) (logical<NImpliesLe (leImpliesLogical<N prBA))
orderIsTransitive : {a b c : } (a <N b) (b <N c) (a <N c)
orderIsTransitive {a} {.(succ (x +N a))} {.(succ (y +N succ (x +N a)))} (le x refl) (le y refl) = le (y +N succ x) (applyEquality succ (additionNIsAssociative y (succ x) a))
<NWellfounded : WellFounded _<N_
<NWellfounded = λ x access (go x)
where
lemma : {a b c : } a <N b b <N succ c a <N c
lemma {a} {b} {c} (le y succYAeqB) (le z zbEqC') = le (y +N z) p
where
zbEqC : z +N b c
zSuccYAEqC : z +N (succ y +N a) c
zSuccYAEqC' : z +N (succ (y +N a)) c
zSuccYAEqC'' : succ (z +N (y +N a)) c
zSuccYAEqC''' : succ ((z +N y) +N a) c
p : succ ((y +N z) +N a) c
p = identityOfIndiscernablesLeft (succ (z +N y) +N a) c (succ (y +N z) +N a) _≡_ zSuccYAEqC''' (applyEquality (λ n succ (n +N a)) (additionNIsCommutative z y))
zSuccYAEqC''' = identityOfIndiscernablesLeft (succ (z +N (y +N a))) c (succ ((z +N y) +N a)) _≡_ zSuccYAEqC'' (equalityCommutative (applyEquality succ (additionNIsAssociative z y a)))
zSuccYAEqC'' = identityOfIndiscernablesLeft (z +N (succ y +N a)) c (succ (z +N (y +N a))) _≡_ zSuccYAEqC' (succExtracts z (y +N a))
zSuccYAEqC' = identityOfIndiscernablesLeft (z +N (succ y +N a)) c (z +N succ (y +N a)) _≡_ zSuccYAEqC (applyEquality (λ r z +N r) refl)
zbEqC = succInjective zbEqC'
zSuccYAEqC = identityOfIndiscernablesLeft (z +N b) c (z +N (succ y +N a)) _≡_ zbEqC (applyEquality (λ r z +N r) (equalityCommutative succYAeqB))
go : n m m <N n Accessible _<N_ m
go zero m (le x ())
go (succ n) zero mLessN = access (λ y ())
go (succ n) (succ m) smLessSN = access (λ o (oLessSM : o <N succ m) go n o (lemma oLessSM smLessSN))
lessImpliesNotEqual : {a b : } (a <N b) a b False
lessImpliesNotEqual {a} {.a} prAB refl = orderIsIrreflexive prAB prAB
-NIsDecreasing : {a b : } (prAB : succ a <N b) subtractionNResult.result (-N (inl prAB)) <N b
-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 {} )
equalityN (succ a) zero = ( BoolFalse , record {} )
equalityN (succ a) (succ b) with equalityN a b
equalityN (succ a) (succ b) | BoolTrue , val = (BoolTrue , applyEquality succ val)
equalityN (succ a) (succ b) | BoolFalse , val = (BoolFalse , record {})
sumZeroImpliesSummandsZero : {a b : } (a +N b zero) ((a zero) && (b zero))
sumZeroImpliesSummandsZero {zero} {zero} pr = record { fst = refl ; snd = refl }
sumZeroImpliesSummandsZero {zero} {(succ b)} pr = record { fst = refl ; snd = pr }
sumZeroImpliesSummandsZero {(succ a)} {zero} ()
sumZeroImpliesSummandsZero {(succ a)} {(succ b)} ()
productWithNonzeroZero : (a b : ) (a *N succ b zero) a zero
productWithNonzeroZero zero b pr = refl
productWithNonzeroZero (succ a) b ()
productOneImpliesOperandsOne : {a b : } (a *N b 1) (a 1) && (b 1)
productOneImpliesOperandsOne {zero} {b} ()
productOneImpliesOperandsOne {succ a} {zero} pr = exFalso absurd''
where
absurd : zero *N (succ a) 1
absurd' : 0 1
absurd'' : False
absurd'' = succIsNonzero (equalityCommutative absurd')
absurd = identityOfIndiscernablesLeft (succ a *N zero) 1 (0 *N succ a) _≡_ pr (productZeroIsZeroRight a)
absurd' = absurd
productOneImpliesOperandsOne {succ a} {succ b} pr = record { fst = r' ; snd = (applyEquality succ (_&&_.fst q)) }
where
p : b +N a *N succ b zero
p = succInjective pr
q : (b zero) && (a *N succ b zero)
q = sumZeroImpliesSummandsZero p
r : a zero
r = productWithNonzeroZero a b (_&&_.snd q)
r' : succ a 1
r' = applyEquality succ r
oneTimesPlusZero : (a : ) a a *N succ zero +N zero
oneTimesPlusZero a = identityOfIndiscernablesRight a (a *N succ zero) (a *N succ zero +N zero) _≡_ (equalityCommutative (productWithOneRight a)) (equalityCommutative (addZeroRight (a *N succ zero)))
cancelInequalityLeft : {a b c : } a *N b <N a *N c b <N c
cancelInequalityLeft {a} {zero} {zero} (le x proof) rewrite (productZeroIsZeroRight a) = exFalso (succIsNonzero {x +N zero} proof)
cancelInequalityLeft {a} {zero} {succ c} pr = succIsPositive c
cancelInequalityLeft {a} {succ b} {zero} (le x proof) rewrite (productZeroIsZeroRight a) = exFalso (succIsNonzero {x +N a *N succ b} proof)
cancelInequalityLeft {a} {succ b} {succ c} pr = succPreservesInequality q'
where
p' : succ b *N a <N succ c *N a
p' = canFlipMultiplicationsInIneq {a} {succ b} {a} {succ c} pr
p'' : b *N a +N a <N succ c *N a
p'' = identityOfIndiscernablesLeft (succ b *N a) (succ c *N a) (b *N a +N a) _<N_ p' (additionNIsCommutative a (b *N a))
p''' : b *N a +N a <N c *N a +N a
p''' = identityOfIndiscernablesRight (b *N a +N a) (succ c *N a) (c *N a +N a) _<N_ p'' (additionNIsCommutative a (c *N a))
p : b *N a <N c *N a
p = subtractionPreservesInequality a p'''
q : a *N b <N a *N c
q = canFlipMultiplicationsInIneq {b} {a} {c} {a} p
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))
cannotAddAndEnlarge' : {a b : } a +N b <N a False
cannotAddAndEnlarge' {a} {zero} pr rewrite addZeroRight a = lessIrreflexive pr
cannotAddAndEnlarge' {a} {succ b} pr rewrite (succExtracts a b) = lessIrreflexive {a} (lessTransitive {a} {succ (a +N b)} {a} (cannotAddAndEnlarge a b) pr)
cannotAddAndEnlarge'' : {a b : } a +N succ b a False
cannotAddAndEnlarge'' {a} {b} pr = naughtE (equalityCommutative inter)
where
inter : succ b 0
inter rewrite additionNIsCommutative a (succ b) = canSubtractFromEqualityRight pr
equivalentSubtraction' : (a b c d : ) (a<c : a <N c) (d<b : d <N b) (subtractionNResult.result (-N {a} {c} (inl a<c))) (subtractionNResult.result (-N {d} {b} (inl d<b))) a +N b c +N d
equivalentSubtraction' a b c d prac prdb eq with -N (inl prac)
equivalentSubtraction' a b c d prac prdb eq | record { result = result ; pr = pr } with -N (inl prdb)
equivalentSubtraction' a b c d prac prdb refl | record { result = .result ; pr = pr1 } | record { result = result ; pr = pr } rewrite (equalityCommutative pr) = go
where
go : a +N (d +N result) c +N d
go rewrite (equalityCommutative pr1) = t
where
t : a +N (d +N result) (a +N result) +N d
t rewrite (additionNIsAssociative a result d) = applyEquality (λ n a +N n) (additionNIsCommutative d result)
lessThanMeansPositiveSubtr : {a b : } (a<b : a <N b) (subtractionNResult.result (-N (inl a<b)) 0) False
lessThanMeansPositiveSubtr {a} {b} a<b pr with -N (inl a<b)
lessThanMeansPositiveSubtr {a} {b} a<b pr | record { result = result ; pr = sub } rewrite pr | addZeroRight a = lessImpliesNotEqual a<b sub
moveOneSubtraction : {a b c : } {a<=b : a ≤N b} (subtractionNResult.result (-N {a} {b} a<=b)) c b a +N c
moveOneSubtraction {a} {b} {zero} {inl a<b} pr rewrite addZeroRight a = exFalso (lessThanMeansPositiveSubtr {a} {b} a<b pr)
moveOneSubtraction {a} {b} {succ c} {inl a<b} pr with -N (inl a<b)
moveOneSubtraction {a} {b} {succ c} {inl a<b} pr | record { result = result ; pr = sub } rewrite pr | sub = refl
moveOneSubtraction {a} {.a} {zero} {inr refl} pr = equalityCommutative (addZeroRight a)
moveOneSubtraction {a} {.a} {succ c} {inr refl} pr = identityOfIndiscernablesRight a (a +N zero) (a +N (succ c)) _≡_ (equalityCommutative (addZeroRight a)) (applyEquality (λ t a +N t) pr')
where
selfSub : (r : ) subtractionNResult.result (-N {r} {r} (inr refl)) zero
selfSub zero = refl
selfSub (succ r) = refl
pr' : 0 succ c
pr' = transitivity (equalityCommutative (selfSub a)) pr
moveOneSubtraction' : {a b c : } {a<=b : a ≤N b} (b a +N c) subtractionNResult.result (-N {a} {b} a<=b) c
moveOneSubtraction' {a} {b} {c} {inl x} pr with -N (inl x)
moveOneSubtraction' {a} {b} {c} {inl x} pr | record { result = result ; pr = pr1 } rewrite pr = canSubtractFromEqualityLeft pr1
moveOneSubtraction' {a} {b} {c} {inr x} pr with -N (inr x)
moveOneSubtraction' {a} {b} {c} {inr x} pr | record { result = result ; pr = pr1 } rewrite pr = canSubtractFromEqualityLeft pr1
equivalentSubtraction : (a b c d : ) (a<c : a <N c) (d<b : d <N b) a +N b c +N d (subtractionNResult.result (-N {a} {c} (inl a<c))) (subtractionNResult.result (-N {d} {b} (inl d<b)))
equivalentSubtraction zero b c d prac (le x proof) eq with (-N (inl (le x proof)))
equivalentSubtraction zero b c d prac (le x proof) eq | record { result = result ; pr = pr } = equalityCommutative p''
where
p : d +N result c +N d
p = transitivity pr eq
p' : d +N result d +N c
p' = transitivity p (additionNIsCommutative c d)
p'' : result c
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'
leLemma : (b c : ) (b ≤N c) (b +N zero ≤N c +N zero)
leLemma b c rewrite addZeroRight c = q
where
q : (b ≤N c) (b +N zero ≤N c)
q rewrite addZeroRight b = refl
lessCast : {a b c : } (pr : a ≤N b) (eq : a c) c ≤N b
lessCast {a} {b} pr eq rewrite eq = pr
lessCast' : {a b c : } (pr : a ≤N b) (eq : b c) a ≤N c
lessCast' {a} {b} pr eq rewrite eq = pr
subtractionCast : {a b c : } {pr : a ≤N b} (eq : a c) (p : subtractionNResult a b pr) Sg (subtractionNResult c b (lessCast pr eq)) (λ res subtractionNResult.result p subtractionNResult.result res)
subtractionCast {a} {b} {c} {a<b} eq subt rewrite eq = (subt , refl)
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))
addToRightWeakInequality zero {b} {.b} (inr refl) = inr (equalityCommutative (addZeroRight b))
addToRightWeakInequality (succ a) {b} {.b} (inr refl) = inl (addingIncreases b a)
addAssocLemma : (a b c : ) (a +N b) +N c (a +N c) +N b
addAssocLemma a b c rewrite (additionNIsAssociative a b c) = g
where
g : a +N (b +N c) (a +N c) +N b
g rewrite (additionNIsAssociative a c b) = applyEquality (λ t a +N t) (additionNIsCommutative b c)
addIntoSubtraction : (a : ) {b c : } (pr : b ≤N c) a +N (subtractionNResult.result (-N {b} {c} pr)) subtractionNResult.result (-N {b} {c +N a} (addToRightWeakInequality a pr))
addIntoSubtraction a {b} {c} pr with (-N {b} {c} pr)
addIntoSubtraction a {b} {c} pr | record { result = c-b ; pr = prc-b } with (-N {b} {c +N a} (addToRightWeakInequality a pr))
addIntoSubtraction a {b} {c} pr | record { result = c-b ; pr = prc-b } | record { result = c+a-b ; pr = prcab } = equalityCommutative g'''
where
g : (b +N c+a-b) +N c-b c +N (a +N c-b)
g rewrite (equalityCommutative (additionNIsAssociative c a c-b)) = applyEquality (λ t t +N c-b) prcab
g' : (b +N c-b) +N c+a-b c +N (a +N c-b)
g' = identityOfIndiscernablesLeft ((b +N c+a-b) +N c-b) (c +N (a +N c-b)) ((b +N c-b) +N c+a-b) _≡_ g (addAssocLemma b c+a-b c-b)
g'' : c +N c+a-b c +N (a +N c-b)
g'' = identityOfIndiscernablesLeft ((b +N c-b) +N c+a-b) (c +N (a +N c-b)) (c +N c+a-b) _≡_ g' (applyEquality (λ i i +N c+a-b) prc-b)
g''' : c+a-b a +N c-b
g''' = canSubtractFromEqualityLeft {c} {c+a-b} {a +N c-b} g''
addStrongInequalities : {a b c d : } (a<b : a <N b) (c<d : c <N d) (a +N c <N b +N d)
addStrongInequalities {zero} {zero} {c} {d} prab prcd = prcd
addStrongInequalities {zero} {succ b} {c} {d} prab prcd rewrite (additionNIsCommutative b d) = orderIsTransitive prcd (cannotAddAndEnlarge d b)
addStrongInequalities {succ a} {zero} {c} {d} (le x ()) prcd
addStrongInequalities {succ a} {succ b} {c} {d} prab prcd = succPreservesInequality (addStrongInequalities (canRemoveSuccFrom<N prab) prcd)
addWeakInequalities : {a b c d : } (a<=b : a ≤N b) (c<=d : c ≤N d) (a +N c) ≤N (b +N d)
addWeakInequalities {a} {b} {c} {d} (inl x) (inl y) = inl (addStrongInequalities x y)
addWeakInequalities {a} {b} {c} {.c} (inl x) (inr refl) = inl (additionPreservesInequality c x)
addWeakInequalities {a} {.a} {c} {d} (inr refl) (inl x) = inl (additionPreservesInequalityOnLeft a x)
addWeakInequalities {a} {.a} {c} {.c} (inr refl) (inr refl) = inr refl
addSubIntoSub : {a b c d : } (a<b : a ≤N b) (c<d : c ≤N d) (subtractionNResult.result (-N {a} {b} a<b)) +N (subtractionNResult.result (-N {c} {d} c<d)) subtractionNResult.result (-N {a +N c} {b +N d} (addWeakInequalities a<b c<d))
addSubIntoSub {a}{b}{c}{d} a<b c<d with (-N {a} {b} a<b)
addSubIntoSub {a} {b} {c} {d} a<b c<d | record { result = b-a ; pr = prb-a } with (-N {c} {d} c<d)
addSubIntoSub {a} {b} {c} {d} a<b c<d | record { result = b-a ; pr = prb-a } | record { result = d-c ; pr = prd-c } with (-N {a +N c} {b +N d} (addWeakInequalities a<b c<d))
addSubIntoSub {a} {b} {c} {d} a<b c<d | record { result = b-a ; pr = prb-a } | record { result = d-c ; pr = prd-c } | record { result = b+d-a-c ; pr = pr } = equalityCommutative r
where
pr' : (a +N c) +N b+d-a-c (a +N b-a) +N d
pr' rewrite prb-a = pr
pr'' : a +N (c +N b+d-a-c) (a +N b-a) +N d
pr'' rewrite (equalityCommutative (additionNIsAssociative a c b+d-a-c)) = pr'
pr''' : a +N (c +N b+d-a-c) a +N (b-a +N d)
pr''' rewrite (equalityCommutative (additionNIsAssociative a b-a d)) = pr''
q : c +N b+d-a-c b-a +N d
q = canSubtractFromEqualityLeft {a} pr'''
q' : c +N b+d-a-c b-a +N (c +N d-c)
q' rewrite prd-c = q
q'' : c +N b+d-a-c (b-a +N c) +N d-c
q'' rewrite additionNIsAssociative b-a c d-c = q'
q''' : c +N b+d-a-c (c +N b-a) +N d-c
q''' rewrite additionNIsCommutative c b-a = q''
q'''' : c +N b+d-a-c c +N (b-a +N d-c)
q'''' rewrite equalityCommutative (additionNIsAssociative c b-a d-c) = q'''
r : b+d-a-c b-a +N d-c
r = canSubtractFromEqualityLeft {c} q''''
subtractProduct : {a b c : } (aPos : 0 <N a) (b<c : b <N c)
(a *N (subtractionNResult.result (-N (inl b<c)))) subtractionNResult.result (-N {a *N b} {a *N c} (inl (lessRespectsMultiplicationLeft b c a b<c aPos)))
subtractProduct {zero} {b} {c} aPos b<c = refl
subtractProduct {succ zero} {b} {c} aPos b<c = s'
where
resbc = -N {b} {c} (inl b<c)
resbc' : Sg (subtractionNResult (b +N 0 *N b) c (lessCast (inl b<c) (equalityCommutative (addZeroRight b)))) ((λ res subtractionNResult.result resbc subtractionNResult.result res))
resbc'' : Sg (subtractionNResult (b +N 0 *N b) (c +N 0 *N c) (lessCast' (lessCast (inl b<c) (equalityCommutative (addZeroRight b))) (equalityCommutative (addZeroRight c)))) (λ res subtractionNResult.result (underlying resbc') subtractionNResult.result res)
g : (rsbc' : Sg (subtractionNResult (b +N 0 *N b) c (lessCast (inl b<c) (equalityCommutative (addZeroRight b)))) (λ res subtractionNResult.result resbc subtractionNResult.result res)) subtractionNResult.result resbc subtractionNResult.result (underlying rsbc')
g' : (rsbc'' : Sg (subtractionNResult (b +N 0 *N b) (c +N 0 *N c) (lessCast' (lessCast (inl b<c) (equalityCommutative (addZeroRight b))) (equalityCommutative (addZeroRight c)))) (λ res subtractionNResult.result (underlying resbc') subtractionNResult.result res)) subtractionNResult.result (underlying resbc') subtractionNResult.result (underlying rsbc'')
q : subtractionNResult.result resbc subtractionNResult.result (underlying resbc'')
r : subtractionNResult.result (underlying resbc'') subtractionNResult.result (-N {b +N 0 *N b} {c +N 0 *N c} (inl (lessRespectsMultiplicationLeft b c 1 b<c aPos)))
s : subtractionNResult.result resbc subtractionNResult.result (-N {b +N 0 *N b} {c +N 0 *N c} (inl (lessRespectsMultiplicationLeft b c 1 b<c aPos)))
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)))
g (a , b) = b
g' (a , b) = b
resbc'' = subtractionCast' (equalityCommutative (addZeroRight c)) (underlying resbc')
q = transitivity {_} {_} {subtractionNResult.result resbc} {subtractionNResult.result (underlying resbc')} {subtractionNResult.result (underlying resbc'')} (g resbc') (g' resbc'')
resbc' = subtractionCast {b} {c} {b +N 0 *N b} (equalityCommutative (addZeroRight b)) resbc
subtractProduct {succ (succ a)} {b} {c} aPos b<c =
let
t : (succ a) *N subtractionNResult.result (-N {b} {c} (inl b<c)) subtractionNResult.result (-N {(succ a) *N b} {(succ a) *N c} (inl (lessRespectsMultiplicationLeft b c (succ a) b<c (succIsPositive a))))
t = subtractProduct {succ a} {b} {c} (succIsPositive a) b<c
in
go t --go t
where
go : (succ a) *N subtractionNResult.result (-N {b} {c} (inl b<c)) subtractionNResult.result (-N {(succ a) *N b} {(succ a) *N c} (inl (lessRespectsMultiplicationLeft b c (succ a) b<c (succIsPositive a)))) (subtractionNResult.result (-N (inl b<c)) +N (subtractionNResult.result (-N (inl b<c)) +N a *N subtractionNResult.result (-N (inl b<c))) subtractionNResult.result (-N (inl (lessRespectsMultiplicationLeft b c (succ (succ a)) b<c aPos))))
go t = transitivity {_} {_} {lhs} {middle2} {rhs} u' v
where
c-b = subtractionNResult.result (-N {b} {c} (inl b<c))
lhs = c-b +N (succ a) *N c-b
middle = subtractionNResult.result (-N {(succ a) *N b} {(succ a) *N c} (inl (lessRespectsMultiplicationLeft b c (succ a) b<c (succIsPositive a))))
middle2 = subtractionNResult.result (-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)))))
rhs = subtractionNResult.result (-N {succ (succ a) *N b} {(succ (succ a)) *N c} (inl (lessRespectsMultiplicationLeft b c (succ (succ a)) b<c aPos)))
lhs' : lhs c-b +N middle
u : c-b +N middle middle2
u' : lhs middle2
v : middle2 rhs
u'' : c-b +N subtractionNResult.result (-N {(succ a) *N b} {(succ a) *N c} (inl (lessRespectsMultiplicationLeft b c (succ a) b<c (succIsPositive a)))) rhs
u'' rewrite equalityCommutative v = u
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)))
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
equalityDecidable (succ a) zero = inr λ t naughtE (equalityCommutative t)
equalityDecidable (succ a) (succ b) with equalityDecidable a b
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)
<NRefl : {a b : } (p1 p2 : a <N b) (p1 p2)
<NRefl {a} {.(succ (p1 +N a))} (le p1 refl) (le p2 proof2) = help p1=p2 proof2
where
p1=p2 : p1 p2
p1=p2 = equalityCommutative (canSubtractFromEqualityRight {p2} {a} {p1} (succInjective proof2))
help : (p1 p2) (pr2 : succ (p2 +N a) succ (p1 +N a)) le p1 refl le p2 pr2
help refl refl = refl
TotalOrder : TotalOrder
PartialOrder._<_ (TotalOrder.order TotalOrder) = _<N_
PartialOrder.irreflexive (TotalOrder.order TotalOrder) = lessIrreflexive
PartialOrder.transitive (TotalOrder.order TotalOrder) = orderIsTransitive
TotalOrder.totality TotalOrder = orderIsTotal
doubleIsAddTwo : (a : ) a +N a 2 *N a
doubleIsAddTwo a rewrite additionNIsCommutative a 0 = refl
productZeroImpliesOperandZero : {a b : } a *N b 0 (a 0) || (b 0)
productZeroImpliesOperandZero {zero} {b} pr = inl refl
productZeroImpliesOperandZero {succ a} {zero} pr = inr refl
productZeroImpliesOperandZero {succ a} {succ b} ()

101
Orders.agda Normal file
View File

@@ -0,0 +1,101 @@
{-# OPTIONS --safe --warning=error #-}
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import LogicalFormulae
open import Functions
module Orders where
record PartialOrder {a b : _} (carrier : Set a) : Set (a lsuc b) where
field
_<_ : Rel {a} {b} carrier
irreflexive : {x : carrier} (x < x) False
transitive : {a b c : carrier} (a < b) (b < c) (a < c)
record TotalOrder {a b : _} (carrier : Set a) : Set (a lsuc b) where
field
order : PartialOrder {a} {b} carrier
_<_ : Rel carrier
_<_ = PartialOrder._<_ order
field
totality : (a b : carrier) ((a < b) || (b < a)) || (a b)
min : carrier carrier carrier
min a b with totality a b
min a b | inl (inl a<b) = a
min a b | inl (inr b<a) = b
min a b | inr a=b = a
max : carrier carrier carrier
max a b with totality a b
max a b | inl (inl a<b) = b
max a b | inl (inr b<a) = a
max a b | inr a=b = b
equivMin : {a b : _} {A : Set a} (order : TotalOrder {a} {b} A) {x y : A} (TotalOrder._<_ order x y) TotalOrder.min order x y x
equivMin order {x} {y} x<y with TotalOrder.totality order x y
equivMin order {x} {y} x<y | inl (inl x₁) = refl
equivMin order {x} {y} x<y | inl (inr y<x) = exFalso (PartialOrder.irreflexive (TotalOrder.order order) (PartialOrder.transitive (TotalOrder.order order) x<y y<x))
equivMin order {x} {y} x<y | inr x=y rewrite x=y = refl
equivMin' : {a b : _} {A : Set a} (order : TotalOrder {a} {b} A) {x y : A} (TotalOrder.min order x y x) (TotalOrder._<_ order x y) || (x y)
equivMin' order {x} {y} minEq with TotalOrder.totality order x y
equivMin' order {x} {y} minEq | inl (inl x<y) = inl x<y
equivMin' order {x} {y} minEq | inl (inr y<x) = exFalso (PartialOrder.irreflexive (TotalOrder.order order) (identityOfIndiscernablesLeft y _ x (TotalOrder._<_ order) y<x minEq))
equivMin' order {x} {y} minEq | inr x=y = inr x=y
minCommutes : {a b : _} {A : Set a} (order : TotalOrder {a} {b} A) (x y : A) (TotalOrder.min order x y) (TotalOrder.min order y x)
minCommutes order x y with TotalOrder.totality order x y
minCommutes order x y | inl (inl x<y) with TotalOrder.totality order y x
minCommutes order x y | inl (inl x<y) | inl (inl y<x) = exFalso (PartialOrder.irreflexive (TotalOrder.order order) (PartialOrder.transitive (TotalOrder.order order) y<x x<y))
minCommutes order x y | inl (inl x<y) | inl (inr x<y') = refl
minCommutes order x y | inl (inl x<y) | inr y=x = equalityCommutative y=x
minCommutes order x y | inl (inr y<x) with TotalOrder.totality order y x
minCommutes order x y | inl (inr y<x) | inl (inl y<x') = refl
minCommutes order x y | inl (inr y<x) | inl (inr x<y) = exFalso (PartialOrder.irreflexive (TotalOrder.order order) (PartialOrder.transitive (TotalOrder.order order) y<x x<y))
minCommutes order x y | inl (inr y<x) | inr y=x = refl
minCommutes order x y | inr x=y with TotalOrder.totality order y x
minCommutes order x y | inr x=y | inl (inl x₁) = x=y
minCommutes order x y | inr x=y | inl (inr x₁) = refl
minCommutes order x y | inr x=y | inr x₁ = x=y
minIdempotent : {a b : _} {A : Set a} (order : TotalOrder {a} {b} A) (x : A) TotalOrder.min order x x x
minIdempotent order x with TotalOrder.totality order x x
minIdempotent order x | inl (inl x₁) = refl
minIdempotent order x | inl (inr x₁) = refl
minIdempotent order x | inr x₁ = refl
swapMin : {a b : _} {A : Set a} {order : TotalOrder {a} {b} A} {x y z : A} (TotalOrder.min order x (TotalOrder.min order y z)) TotalOrder.min order y (TotalOrder.min order x z)
swapMin {order = order} {x} {y} {z} with TotalOrder.totality order y z
swapMin {order = order} {x} {y} {z} | inl (inl y<z) with TotalOrder.totality order x z
swapMin {order = order} {x} {y} {z} | inl (inl y<z) | inl (inl x<z) = minCommutes order x y
swapMin {order = order} {x} {y} {z} | inl (inl y<z) | inl (inr z<x) with TotalOrder.totality order x y
swapMin {order = order} {x} {y} {z} | inl (inl y<z) | inl (inr z<x) | inl (inl x<y) = exFalso (PartialOrder.irreflexive (TotalOrder.order order) (PartialOrder.transitive (TotalOrder.order order) y<z (PartialOrder.transitive (TotalOrder.order order) z<x x<y)))
swapMin {order = order} {x} {y} {z} | inl (inl y<z) | inl (inr z<x) | inl (inr y<x) = equalityCommutative (equivMin order y<z)
swapMin {order = order} {x} {y} {z} | inl (inl y<z) | inl (inr z<x) | inr x=y rewrite x=y = equalityCommutative (equivMin order y<z)
swapMin {order = order} {x} {y} {z} | inl (inl y<z) | inr x=z = minCommutes order x y
swapMin {order = order} {x} {y} {z} | inl (inr z<y) with TotalOrder.totality order x z
swapMin {order = order} {x} {y} {z} | inl (inr z<y) | inl (inl x<z) rewrite minCommutes order y x = equalityCommutative (equivMin order (PartialOrder.transitive (TotalOrder.order order) x<z z<y))
swapMin {order = order} {x} {y} {z} | inl (inr z<y) | inl (inr z<x) with TotalOrder.totality order y z
swapMin {order = order} {x} {y} {z} | inl (inr z<y) | inl (inr z<x) | inl (inl y<z) = exFalso (PartialOrder.irreflexive (TotalOrder.order order) (PartialOrder.transitive (TotalOrder.order order) z<y y<z))
swapMin {order = order} {x} {y} {z} | inl (inr z<y) | inl (inr z<x) | inl (inr z<y') = refl
swapMin {order = order} {x} {y} {z} | inl (inr z<y) | inl (inr z<x) | inr y=z = equalityCommutative y=z
swapMin {order = order} {x} {y} {z} | inl (inr z<y) | inr x=z rewrite x=z | minCommutes order y z = equalityCommutative (equivMin order z<y)
swapMin {order = order} {x} {y} {z} | inr y=z with TotalOrder.totality order x z
swapMin {order = order} {x} {y} {z} | inr y=z | inl (inl x<z) = minCommutes order x y
swapMin {order = order} {x} {y} {z} | inr y=z | inl (inr z<x) rewrite y=z | minIdempotent order z | minCommutes order x z = equivMin order z<x
swapMin {order = order} {x} {y} {z} | inr y=z | inr x=z = minCommutes order x y
minMin : {a b : _} {A : Set a} (order : TotalOrder {a} {b} A) {x y : A} (TotalOrder.min order x (TotalOrder.min order x y)) TotalOrder.min order x y
minMin order {x} {y} with TotalOrder.totality order x y
minMin order {x} {y} | inl (inl x<y) = minIdempotent order x
minMin order {x} {y} | inl (inr y<x) with TotalOrder.totality order x y
minMin order {x} {y} | inl (inr y<x) | inl (inl x<y) = exFalso (PartialOrder.irreflexive (TotalOrder.order order) (PartialOrder.transitive (TotalOrder.order order) y<x x<y))
minMin order {x} {y} | inl (inr y<x) | inl (inr y<x') = refl
minMin order {x} {y} | inl (inr y<x) | inr x=y = x=y
minMin order {x} {y} | inr x=y = minIdempotent order x
minFromBoth : {a b : _} {A : Set a} {order : TotalOrder {a} {b} A} {min x y : A} (TotalOrder._<_ order min x) (TotalOrder._<_ order min y) (TotalOrder._<_ order min (TotalOrder.min order x y))
minFromBoth {a} {order = order} {x = x} {y} prX prY with TotalOrder.totality order x y
minFromBoth {a} prX prY | inl (inl x<y) = prX
minFromBoth {a} prX prY | inl (inr y<x) = prY
minFromBoth {a} prX prY | inr x=y = prX

802
PrimeNumbers.agda Normal file
View File

@@ -0,0 +1,802 @@
{-# OPTIONS --warning=error --safe #-}
open import LogicalFormulae
open import Naturals
open import WellFoundedInduction
open import KeyValue
open import KeyValueWithDomain
open import Orders
open import Vectors
open import Maybe
module PrimeNumbers where
record divisionAlgResult (a : ) (b : ) : Set where
field
quot :
rem :
pr : a *N quot +N rem b
remIsSmall : (rem <N a) || (a 0)
quotSmall : (0 <N a) || ((0 a) && (quot 0))
divAlgLessLemma : (a b : ) (0 <N a) (r : divisionAlgResult a b) (divisionAlgResult.quot r 0) || (divisionAlgResult.rem r <N b)
divAlgLessLemma zero b pr _ = exFalso (lessIrreflexive pr)
divAlgLessLemma (succ a) b _ record { quot = zero ; rem = a%b ; pr = pr ; remIsSmall = remIsSmall } = inl refl
divAlgLessLemma (succ a) b _ record { quot = (succ a/b) ; rem = a%b ; pr = pr ; remIsSmall = remIsSmall } = inr record { x = a/b +N a *N succ (a/b) ; proof = pr }
modUniqueLemma : {rem1 rem2 a : } (quot1 quot2 : ) rem1 <N a rem2 <N a a *N quot1 +N rem1 a *N quot2 +N rem2 rem1 rem2
modUniqueLemma {rem1} {rem2} {a} zero zero rem1<a rem2<a pr rewrite productZeroIsZeroRight a = pr
modUniqueLemma {rem1} {rem2} {a} zero (succ quot2) rem1<a rem2<a pr rewrite productZeroIsZeroRight a | pr | multiplicationNIsCommutative a (succ quot2) | additionNIsAssociative a (quot2 *N a) rem2 = exFalso (cannotAddAndEnlarge' {a} {quot2 *N a +N rem2} rem1<a)
modUniqueLemma {rem1} {rem2} {a} (succ quot1) zero rem1<a rem2<a pr rewrite productZeroIsZeroRight a | equalityCommutative pr | multiplicationNIsCommutative a (succ quot1) | additionNIsAssociative a (quot1 *N a) rem1 = exFalso (cannotAddAndEnlarge' {a} {quot1 *N a +N rem1} rem2<a)
modUniqueLemma {rem1} {rem2} {a} (succ quot1) (succ quot2) rem1<a rem2<a pr rewrite multiplicationNIsCommutative a (succ quot1) | multiplicationNIsCommutative a (succ quot2) | additionNIsAssociative a (quot1 *N a) rem1 | additionNIsAssociative a (quot2 *N a) rem2 = modUniqueLemma {rem1} {rem2} {a} quot1 quot2 rem1<a rem2<a (go {a}{quot1}{rem1}{quot2}{rem2} pr)
where
go : {a quot1 rem1 quot2 rem2 : } (a +N (quot1 *N a +N rem1) a +N (quot2 *N a +N rem2)) a *N quot1 +N rem1 a *N quot2 +N rem2
go {a} {quot1} {rem1} {quot2} {rem2} pr rewrite multiplicationNIsCommutative quot1 a | multiplicationNIsCommutative quot2 a = canSubtractFromEqualityLeft {a} pr
modIsUnique : {a b : } (div1 div2 : divisionAlgResult a b) divisionAlgResult.rem div1 divisionAlgResult.rem div2
modIsUnique {zero} {b} record { quot = quot1 ; rem = rem1 ; pr = pr1 ; remIsSmall = remIsSmall1 } record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = remIsSmall } = transitivity pr1 (equalityCommutative pr)
modIsUnique {succ a} {b} record { quot = quot1 ; rem = rem1 ; pr = pr1 ; remIsSmall = (inl y) } record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = (inl x) } = modUniqueLemma {rem1} {rem} {succ a} quot1 quot y x (transitivity pr1 (equalityCommutative pr))
modIsUnique {succ a} {b} record { quot = quot1 ; rem = rem1 ; pr = pr1 ; remIsSmall = (inr ()) } record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = (inl x) }
modIsUnique {succ a} {b} record { quot = quot1 ; rem = rem1 ; pr = pr1 ; remIsSmall = remIsSmall1 } record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = (inr ()) }
transferAddition : (a b c : ) (a +N b) +N c (a +N c) +N b
transferAddition a b c rewrite (additionNIsAssociative a b c) = p a b c
where
p : (a b c : ) a +N (b +N c) (a +N c) +N b
p a b c rewrite (additionNIsCommutative b c) = equalityCommutative (additionNIsAssociative a c b)
divisionAlgLemma : (x b : ) x *N zero +N b b
divisionAlgLemma x b rewrite (productZeroIsZeroRight x) = refl
divisionAlgLemma2 : (x b : ) (x b) x *N succ zero +N zero b
divisionAlgLemma2 x b pr rewrite (productWithOneRight x) = equalityCommutative (transitivity (equalityCommutative pr) (equalityCommutative (addZeroRight x)))
divisionAlgLemma3 : {a x : } (p : succ a <N succ x) (subtractionNResult.result (-N (inl p))) <N (succ x)
divisionAlgLemma3 {a} {x} p = -NIsDecreasing {a} {succ x} p
divisionAlgLemma4 : (p a q : ) ((p +N a *N p) +N q) +N succ a succ ((p +N a *N succ p) +N q)
divisionAlgLemma4 p a q = ans
where
r : ((p +N a *N p) +N q) +N succ a succ (((p +N a *N p) +N q) +N a)
ans : ((p +N a *N p) +N q) +N succ a succ ((p +N a *N succ p) +N q)
s : ((p +N a *N p) +N q) +N a (p +N a *N succ p) +N q
t : (p +N a *N p) +N a p +N a *N succ p
s = transitivity (transferAddition (p +N a *N p) q a) (applyEquality (λ i i +N q) t)
ans = identityOfIndiscernablesRight (((p +N a *N p) +N q) +N succ a) (succ (((p +N a *N p) +N q) +N a)) (succ ((p +N a *N succ p) +N q)) _≡_ r (applyEquality succ s)
r = succExtracts ((p +N a *N p) +N q) a
t = transitivity (additionNIsAssociative p (a *N p) a) (applyEquality (λ n p +N n) (equalityCommutative (aSucB a p)))
divisionAlg : (a : ) (b : ) divisionAlgResult a b
divisionAlg zero = λ b record { quot = zero ; rem = b ; pr = refl ; remIsSmall = inr refl ; quotSmall = inr (record { fst = refl ; snd = refl }) }
divisionAlg (succ a) = rec <NWellfounded (λ n divisionAlgResult (succ a) n) go
where
go : (x : ) (indHyp : (y : ) (y<x : y <N x) divisionAlgResult (succ a) y)
divisionAlgResult (succ a) x
go zero prop = record { quot = zero ; rem = zero ; pr = divisionAlgLemma (succ a) zero ; remIsSmall = inl (succIsPositive a) ; quotSmall = inl (succIsPositive a) }
go (succ x) indHyp with orderIsTotal (succ a) (succ x)
go (succ x) indHyp | inl (inl sa<sx) with indHyp (subtractionNResult.result (-N (inl sa<sx))) (divisionAlgLemma3 sa<sx)
... | record { quot = prevQuot ; rem = prevRem ; pr = prevPr ; remIsSmall = smallRem } = p
where
p : divisionAlgResult (succ a) (succ x)
addedA : (succ a *N prevQuot +N prevRem) +N (succ a) subtractionNResult.result (-N (inl sa<sx)) +N (succ a)
addedA' : (succ a *N prevQuot +N prevRem) +N succ a succ x
addedA'' : (succ a *N succ prevQuot) +N prevRem succ x
addedA''' : succ ((prevQuot +N a *N succ prevQuot) +N prevRem) succ x
addedA''' = identityOfIndiscernablesLeft ((succ a *N succ prevQuot) +N prevRem) (succ x) (succ ((prevQuot +N a *N succ prevQuot) +N prevRem)) _≡_ addedA'' refl
p = record { quot = succ prevQuot ; rem = prevRem ; pr = addedA''' ; remIsSmall = smallRem ; quotSmall = inl (succIsPositive a) }
addedA = applyEquality (λ n n +N succ a) prevPr
addedA' = identityOfIndiscernablesRight ((succ a *N prevQuot +N prevRem) +N succ a) (subtractionNResult.result (-N (inl sa<sx)) +N (succ a)) (succ x) _≡_ addedA (addMinus {succ a} {succ x} (inl sa<sx))
addedA'' = identityOfIndiscernablesLeft ((succ a *N prevQuot +N prevRem) +N succ a) (succ x) ((succ a *N succ prevQuot) +N prevRem) _≡_ addedA' (divisionAlgLemma4 prevQuot a prevRem)
go (succ x) indHyp | inr (sa=sx) = record { quot = succ zero ; rem = zero ; pr = divisionAlgLemma2 (succ a) (succ x) sa=sx ; remIsSmall = inl (succIsPositive a) ; quotSmall = inl (succIsPositive a) }
go (succ x) indHyp | inl (inr (sx<sa)) = record { quot = zero ; rem = succ x ; pr = divisionAlgLemma (succ a) (succ x) ; remIsSmall = inl sx<sa ; quotSmall = inl (succIsPositive a) }
data __ : Set where
divides : {a b : } (res : divisionAlgResult a b) divisionAlgResult.rem res zero a b
dividesEqualityLemma'' : {a b : } (quot1 quot2 : ) (quot1 quot2) (rem : ) (pr1 : (quot1 +N a *N quot1) +N rem b) (pr2 : (quot2 +N a *N quot2) +N rem b) (y : rem <N succ a) (x1 : zero <N succ a) record { quot = quot1 ; rem = rem ; pr = pr1 ; remIsSmall = inl y ; quotSmall = inl x1 } record { quot = quot2 ; rem = rem ; pr = pr2 ; remIsSmall = inl y ; quotSmall = inl x1}
dividesEqualityLemma'' {a} {b} q1 .q1 refl rem pr1 pr2 y x1 rewrite reflRefl pr1 pr2 = refl
dividesEqualityLemma' : {a b : } (quot1 quot2 : ) (rem : ) (pr1 : (quot1 +N a *N quot1) +N rem b) (pr2 : (quot2 +N a *N quot2) +N rem b) (y : rem <N succ a) (y2 : rem <N succ a) (x1 : zero <N succ a) record { quot = quot1 ; rem = rem ; pr = pr1 ; remIsSmall = inl y ; quotSmall = inl x1 } record { quot = quot2 ; rem = rem ; pr = pr2 ; remIsSmall = inl y2 ; quotSmall = inl x1}
dividesEqualityLemma' {a} {b} quot1 quot2 rem pr1 pr2 y y2 x1 rewrite <NRefl y2 y = dividesEqualityLemma'' quot1 quot2 (equalityCommutative lemm'') rem pr1 pr2 y x1
where
lemm : (succ a) *N quot2 +N rem (succ a) *N quot1 +N rem
lemm rewrite equalityCommutative pr1 = pr2
lemm' : (succ a) *N quot2 (succ a) *N quot1
lemm' = canSubtractFromEqualityRight lemm
lemm'' : quot2 quot1
lemm'' = productCancelsLeft (succ a) quot2 quot1 (succIsPositive a) lemm'
dividesEqualityLemma : {a b : } (quot1 quot2 : ) (rem1 rem2 : ) (pr1 : (quot1 +N a *N quot1) +N rem1 b) (pr2 : (quot2 +N a *N quot2) +N rem2 b) (rem1 rem2) (y : rem1 <N succ a) (y2 : rem2 <N succ a) (x1 : zero <N succ a) record { quot = quot1 ; rem = rem1 ; pr = pr1 ; remIsSmall = inl y ; quotSmall = inl x1 } record { quot = quot2 ; rem = rem2 ; pr = pr2 ; remIsSmall = inl y2 ; quotSmall = inl x1}
dividesEqualityLemma {a} {b} quot1 quot2 rem1 rem2 pr1 pr2 remEq y y2 x1 rewrite remEq = dividesEqualityLemma' quot1 quot2 rem2 pr1 pr2 y y2 x1
dividesEqualityPr : {a b : } (res1 res2 : divisionAlgResult a b) res1 res2
dividesEqualityPr {zero} {b} record { quot = quot1 ; rem = rem1 ; pr = pr1 ; remIsSmall = remIsSmall1 ; quotSmall = (inl (le x ())) } record { quot = quot2 ; rem = rem2 ; pr = pr2 ; remIsSmall = remIsSmall2 ; quotSmall = quotSmall2 }
dividesEqualityPr {zero} {b} record { quot = quot1 ; rem = rem1 ; pr = pr1 ; remIsSmall = remIsSmall1 ; quotSmall = (inr (fst ,, snd)) } record { quot = quot2 ; rem = rem2 ; pr = pr2 ; remIsSmall = remIsSmall2 ; quotSmall = (inl (le x ())) }
dividesEqualityPr {zero} {b} record { quot = .0 ; rem = rem1 ; pr = pr1 ; remIsSmall = (inl (le x ())) ; quotSmall = (inr (refl ,, refl)) } record { quot = .0 ; rem = rem2 ; pr = pr2 ; remIsSmall = remIsSmall2 ; quotSmall = (inr (refl ,, refl)) }
dividesEqualityPr {zero} {b} record { quot = .0 ; rem = rem1 ; pr = pr1 ; remIsSmall = (inr refl) ; quotSmall = (inr (refl ,, refl)) } record { quot = .0 ; rem = rem2 ; pr = pr2 ; remIsSmall = (inl (le x ())) ; quotSmall = (inr (refl ,, refl)) }
dividesEqualityPr {zero} {.rem1} record { quot = .0 ; rem = rem1 ; pr = refl ; remIsSmall = (inr refl) ; quotSmall = (inr (refl ,, refl)) } record { quot = .0 ; rem = .rem1 ; pr = refl ; remIsSmall = (inr refl) ; quotSmall = (inr (refl ,, refl)) } = refl
dividesEqualityPr {succ a} {b} record { quot = quot1 ; rem = rem1 ; pr = pr1 ; remIsSmall = (inl y) ; quotSmall = (inl x) } record { quot = quot2 ; rem = rem2 ; pr = pr2 ; remIsSmall = (inl y2) ; quotSmall = (inl x1) } rewrite <NRefl x x1 = dividesEqualityLemma quot1 quot2 rem1 rem2 pr1 pr2 remsEqual y y2 x1
where
remsEqual : rem1 rem2
remsEqual = modIsUnique (record { quot = quot1 ; rem = rem1 ; pr = pr1 ; quotSmall = (inl x) ; remIsSmall = (inl y) }) (record { quot = quot2 ; rem = rem2 ; pr = pr2 ; remIsSmall = (inl y2) ; quotSmall = (inl x1) })
dividesEqualityPr {succ a} {b} record { quot = quot1 ; rem = rem1 ; pr = pr1 ; remIsSmall = (inl y) ; quotSmall = (inl x) } record { quot = quot2 ; rem = rem2 ; pr = pr2 ; remIsSmall = (inr ()) ; quotSmall = (inl x1) }
dividesEqualityPr {succ a} {b} record { quot = quot1 ; rem = rem1 ; pr = pr1 ; remIsSmall = (inr ()) ; quotSmall = (inl x) } record { quot = quot2 ; rem = rem2 ; pr = pr2 ; remIsSmall = remIsSmall2 ; quotSmall = (inl x1) }
dividesEqualityPr {succ a} {b} record { quot = quot1 ; rem = rem1 ; pr = pr1 ; remIsSmall = remIsSmall1 ; quotSmall = (inl x) } record { quot = quot2 ; rem = rem2 ; pr = pr2 ; remIsSmall = remIsSmall2 ; quotSmall = (inr (() ,, snd)) }
dividesEqualityPr {succ a} {b} record { quot = quot1 ; rem = rem1 ; pr = pr1 ; remIsSmall = remIsSmall1 ; quotSmall = (inr (() ,, snd)) } record { quot = quot2 ; rem = rem2 ; pr = pr2 ; remIsSmall = remIsSmall2 ; quotSmall = quotSmall2 }
dividesEquality : {a b : } (res1 res2 : a b) res1 res2
dividesEquality (divides res1 x1) (divides res2 x2) rewrite dividesEqualityPr res1 res2 = applyEquality (λ i divides res2 i) (reflRefl x1 x2)
data notDiv : Set where
doesNotDivide : {a b : } (res : divisionAlgResult a b) 0 <N divisionAlgResult.rem res notDiv a b
zeroDividesNothing : (a : ) zero succ a False
zeroDividesNothing a (divides record { quot = quot ; rem = rem ; pr = pr } x) = naughtE p
where
p : zero succ a
p = transitivity (equalityCommutative x) pr
twoDividesFour : succ (succ zero) succ (succ (succ (succ zero)))
twoDividesFour = divides {(succ (succ zero))} {succ (succ (succ (succ zero)))} (record { quot = succ (succ zero) ; rem = zero ; pr = refl ; remIsSmall = inl (succIsPositive 1) ; quotSmall = inl (succIsPositive 1) }) refl
record Prime (p : ) : Set where
field
p>1 : 1 <N p
pr : forall {i : } i p i <N p zero <N i i (succ zero)
record Composite (n : ) : Set where
field
n>1 : 1 <N n
divisor :
dividesN : divisor n
divisorLessN : divisor <N n
divisorNot1 : 1 <N divisor
divisorPrime : Prime divisor
noSmallerDivisors : i i <N divisor 1 <N i i n False
notBothPrimeAndComposite : {n : } Composite n Prime n False
notBothPrimeAndComposite {n} record { n>1 = n>1 ; divisor = divisor ; dividesN = dividesN ; divisorLessN = divisorLessN ; divisorNot1 = divisorNot1 } record { p>1 = p>1 ; pr = pr } = lessImpliesNotEqual divisorNot1 (equalityCommutative div=1)
where
div=1 : divisor 1
div=1 = pr {divisor} dividesN divisorLessN (orderIsTransitive (succIsPositive 0) divisorNot1)
zeroIsNotPrime : Prime 0 False
zeroIsNotPrime record { p>1 = p>1 ; pr = pr } = zeroNeverGreater p>1
oneIsNotPrime : Prime 1 False
oneIsNotPrime record { p>1 = (le x proof) ; pr = pr } = naughtE (equalityCommutative absurd')
where
absurd : x +N 1 0
absurd = succInjective proof
absurd' : succ x 0
absurd' rewrite additionNIsCommutative 1 x = absurd
twoIsPrime : Prime 2
Prime.p>1 twoIsPrime = succPreservesInequality (succIsPositive 0)
Prime.pr twoIsPrime {i} i|2 i<2 0<i with orderIsTotal i (succ (succ zero))
Prime.pr twoIsPrime {zero} i|2 i<2 (le x ()) | order
Prime.pr twoIsPrime {succ zero} i|2 i<2 0<i | order = refl
Prime.pr twoIsPrime {succ (succ zero)} i|2 i<2 0<i | order = exFalso (lessImpliesNotEqual {2} i<2 refl)
Prime.pr twoIsPrime {succ (succ (succ i))} i|2 i<2 0<i | inl (inl x) = exFalso (orderIsIrreflexive i<2 (succPreservesInequality (succPreservesInequality (succIsPositive i))))
Prime.pr twoIsPrime {succ (succ (succ i))} i|2 i<2 0<i | inl (inr twoLessThree) = exFalso (orderIsIrreflexive twoLessThree i<2)
Prime.pr twoIsPrime {succ (succ (succ i))} i|2 i<2 0<i | inr ()
compositeImpliesNotPrime : (m p : ) (succ zero <N m) (m <N p) (m p) Prime p False
compositeImpliesNotPrime zero p (le x ()) _ mDivP pPrime
compositeImpliesNotPrime (succ zero) p mLarge _ mDivP pPrime = lessImpliesNotEqual {succ zero} {succ zero} mLarge refl
compositeImpliesNotPrime (succ (succ m)) zero _ _ mDivP ()
compositeImpliesNotPrime (succ (succ m)) (succ zero) _ _ mDivP primeP = exFalso (oneIsNotPrime primeP)
compositeImpliesNotPrime (succ (succ m)) (succ (succ p)) _ mLessP mDivP pPrime = false
where
r = succ (succ m)
q = succ (succ p)
rEqOne : r succ zero
rEqOne = (Prime.pr pPrime) {r} mDivP mLessP (succIsPositive (succ m))
false : False
false = succIsNonzero (succInjective rEqOne)
fourIsNotPrime : Prime 4 False
fourIsNotPrime = compositeImpliesNotPrime (succ (succ zero)) (succ (succ (succ (succ zero)))) (le zero refl) (le (succ zero) refl) twoDividesFour
record hcfData (a b : ) : Set where
field
c :
c|a : c a
c|b : c b
hcf : x x a x b x c
record Coprime (a : ) (b : ) : Set where
field
hcf : hcfData a b
hcfNot1 : 1 <N hcfData.c hcf
record numberLessThan (n : ) : Set where
field
a :
a<n : a <N n
upper :
upper = n
numberLessThanEquality : {n : } (a b : numberLessThan n) (numberLessThan.a a numberLessThan.a b) a b
numberLessThanEquality record { a = a ; a<n = a<n } record { a = b ; a<n = b<n } pr rewrite pr | <NRefl b<n a<n = refl
numberLessThanOrder : (n : ) TotalOrder (numberLessThan n)
PartialOrder._<_ (TotalOrder.order (numberLessThanOrder n)) = λ a b (numberLessThan.a a) <N numberLessThan.a b
PartialOrder.irreflexive (TotalOrder.order (numberLessThanOrder n)) pr = lessIrreflexive pr
PartialOrder.transitive (TotalOrder.order (numberLessThanOrder n)) pr1 pr2 = orderIsTransitive pr1 pr2
TotalOrder.totality (numberLessThanOrder n) a b with TotalOrder.totality TotalOrder (numberLessThan.a a) (numberLessThan.a b)
TotalOrder.totality (numberLessThanOrder n) a b | inl (inl x) = inl (inl x)
TotalOrder.totality (numberLessThanOrder n) a b | inl (inr x) = inl (inr x)
TotalOrder.totality (numberLessThanOrder n) a b | inr x rewrite x = inr (numberLessThanEquality a b x)
numberLessThanInject : {newMax : } (max : ) (n : numberLessThan max) (max <N newMax) (numberLessThan newMax)
numberLessThanInject max record { a = n ; a<n = n<max } max<newMax = record { a = n ; a<n = PartialOrder.transitive (TotalOrder.order TotalOrder) n<max max<newMax }
numberLessThanInjectComp : {max : } (a b : ) (i : numberLessThan b) (pr : b <N a) (pr2 : a <N max) numberLessThanInject {max} a (numberLessThanInject {a} b i pr) pr2 numberLessThanInject {max} b i (PartialOrder.transitive (TotalOrder.order TotalOrder) pr pr2)
numberLessThanInjectComp {max} a b record { a = i ; a<n = i<max } b<a a<max = numberLessThanEquality _ _ refl
allNumbersLessThanDescending : (n : ) Vec (numberLessThan n) n
allNumbersLessThanDescending zero = []
allNumbersLessThanDescending (succ n) = record { a = n ; a<n = le zero refl } ,- vecMap (λ i numberLessThanInject {succ n} (numberLessThan.upper i) i (le zero refl)) (allNumbersLessThanDescending n)
allNumbersLessThan : (n : ) Vec (numberLessThan n) n
allNumbersLessThan n = vecRev (allNumbersLessThanDescending n)
record extensionalHCF (a b : ) : Set where
field
c :
c|a : c a
c|b : c b
field
zeroCase : ((a 0) & (b 0) & (c 0)) || ((0 <N a) || (0 <N b))
hcfExtension : MapWithDomain (numberLessThan c) (Sg (numberLessThan c) (λ i ((notDiv (numberLessThan.a i) a) || (notDiv (numberLessThan.a i) b)) || ((numberLessThan.a i) a & (numberLessThan.a i) b & (numberLessThan.a i) c))) (numberLessThanOrder c)
hcfExtensionIsRightLength : vecLen (MapWithDomain.domain hcfExtension) c
{-
hcfsContains : {a b r : } → (hcf : extensionalHCF a b) → (r<hcf : r <N extensionalHCF.c hcf) → vecContains (MapWithDomain.domain (extensionalHCF.hcfExtension hcf)) record { a = r ; a<n = r<hcf }
hcfsContains = {!!}
hcfsEquivalent : {a b : } → hcfData a b → extensionalHCF a b
hcfsEquivalent {a} {b} record { c = c ; c|a = c|a ; c|b = c|b ; hcf = hcf } = record { c = c ; c|a = c|a ; c|b = c|b ; hcfExtension = hcfsMap ; hcfExtensionIsRightLength = {!!} ; zeroCase = {!!} }
where
pair : Set
pair = (Sg (numberLessThan c) (λ i → ((notDiv (numberLessThan.a i) a) || (notDiv (numberLessThan.a i) b)) || ((numberLessThan.a i) a & (numberLessThan.a i) b & (numberLessThan.a i) c)))
allHcfs : Map (numberLessThan c) pair (numberLessThanOrder c)
allHcfs = {!!}
hcfsMap : MapWithDomain (numberLessThan c) (Sg (numberLessThan c) (λ i → ((notDiv (numberLessThan.a i) a) || (notDiv (numberLessThan.a i) b)) || ((numberLessThan.a i) a & (numberLessThan.a i) b & (numberLessThan.a i) c))) (numberLessThanOrder c)
hcfsMap = {!!}
-}
positiveTimes : {a b : } (succ a *N succ b <N succ a) False
positiveTimes {a} {b} pr = zeroNeverGreater f'
where
g : succ a *N succ b <N succ a *N succ 0
g rewrite multiplicationNIsCommutative a 1 | additionNIsCommutative a 0 = pr
f : succ b <N succ 0
f = cancelInequalityLeft {succ a} {succ b} g
f' : b <N 0
f' = canRemoveSuccFrom<N f
biggerThanCantDivideLemma : {a b : } (a <N b) (b a) a 0
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 = (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)
biggerThanCantDivide {a} {b} x pr x|a x|b with TotalOrder.totality TotalOrder a b
biggerThanCantDivide {a} {b} x pr x|a x|b | inl (inl a<b) = exFalso (zeroNeverGreater f')
where
f : b 0
f = biggerThanCantDivideLemma pr x|b
f' : a <N 0
f' rewrite equalityCommutative f = a<b
biggerThanCantDivide {a} {b} x pr x|a x|b | inl (inr b<a) = exFalso (zeroNeverGreater f')
where
f : a 0
f = biggerThanCantDivideLemma pr x|a
f' : b <N 0
f' rewrite equalityCommutative f = b<a
biggerThanCantDivide {a} {b} x pr x|a x|b | inr a=b = (transitivity a=b f ,, f)
where
f : b 0
f = biggerThanCantDivideLemma pr x|b
aDivA : (a : ) a a
aDivA zero = divides (record { quot = 0 ; rem = 0 ; pr = equalityCommutative (oneTimesPlusZero zero) ; remIsSmall = inr refl ; quotSmall = inr (refl ,, refl) }) refl
aDivA (succ a) = divides (record { quot = 1 ; rem = 0 ; pr = equalityCommutative (oneTimesPlusZero (succ a)) ; remIsSmall = inl (succIsPositive a) ; quotSmall = inl (succIsPositive a) }) refl
aDivZero : (a : ) a zero
aDivZero zero = aDivA zero
aDivZero (succ a) = divides (record { quot = zero ; rem = zero ; pr = lemma (succ a) ; remIsSmall = inl (succIsPositive a) ; quotSmall = inl (succIsPositive a) }) refl
where
lemma : (b : ) b *N zero +N zero zero
lemma b rewrite (addZeroRight (b *N zero)) = productZeroIsZeroRight b
maxDivides : (a b : ) ((TotalOrder.max TotalOrder a b) a) (TotalOrder.max TotalOrder a b) b (((a 0) && (0 <N b)) || ((b 0) && (0 <N a))) || (a b)
maxDivides a b max|a max|b with TotalOrder.totality TotalOrder a b
maxDivides a b max|a max|b | inl (inl a<b) = inl (inl (record { fst = gg ; snd = identityOfIndiscernablesLeft _ _ _ _<N_ a<b gg}))
where
gg : a 0
gg = biggerThanCantDivideLemma {a} {b} a<b max|a
maxDivides a b max|a max|b | inl (inr b<a) = inl (inr (record { fst = gg ; snd = identityOfIndiscernablesLeft _ _ _ _<N_ b<a gg }))
where
gg : b 0
gg = biggerThanCantDivideLemma b<a max|b
maxDivides a .a a|max b|max | inr refl = inr refl
{-
hcfsEquivalent' : {a b : } → extensionalHCF a b → hcfData a b
hcfData.c (hcfsEquivalent' {a} {b} record { c = c ; c|a = c|a ; c|b = c|b ; hcfExtension = hcfExtension }) = c
hcfData.c|a (hcfsEquivalent' {a} {b} record { c = c ; c|a = c|a ; c|b = c|b ; hcfExtension = hcfExtension }) = c|a
hcfData.c|b (hcfsEquivalent' {a} {b} record { c = c ; c|a = c|a ; c|b = c|b ; hcfExtension = hcfExtension }) = c|b
hcfData.hcf (hcfsEquivalent' {a} {b} record { c = c ; c|a = c|a ; c|b = c|b ; hcfExtension = hcfExtension }) x x|a x|b with orderIsTotal x c
hcfData.hcf (hcfsEquivalent' {a} {b} record { c = c ; c|a = c|a ; c|b = c|b ; zeroCase = zeroCase ; hcfExtension = hcfExtension ; hcfExtensionIsRightLength = hIRL }) x x|a x|b | inl (inl x<c) = {!!}
where
xLess : numberLessThan c
xLess = record { a = x ; a<n = x<c }
pair : Set
pair = Sg (numberLessThan c) (λ i → (notDiv (numberLessThan.a i) a || notDiv (numberLessThan.a i) b) || (numberLessThan.a i a) & numberLessThan.a i b & (numberLessThan.a i c))
pr : Sg pair λ p → (lookup (MapWithDomain.map hcfExtension) xLess ≡ {!!})
pr = MapWithDomain.lookup' hcfExtension xLess (hcfsContains {a} {b} {x} (record { c = c ; c|a = c|a ; c|b = c|b ; zeroCase = zeroCase ; hcfExtension = hcfExtension ; hcfExtensionIsRightLength = hIRL }) x<c)
hcfData.hcf (hcfsEquivalent' {a} {b} record { c = c ; c|a = c|a ; c|b = c|b ; zeroCase = _ ; hcfExtension = hcfExtension ; hcfExtensionIsRightLength = _ }) x x|a x|b | inl (inr c<x) = {!!}
hcfData.hcf (hcfsEquivalent' {a} {b} record { c = c ; c|a = c|a ; c|b = c|b ; zeroCase = _ ; hcfExtension = hcfExtension ; hcfExtensionIsRightLength = _ }) x x|a x|b | inr x=c rewrite x=c = aDivA c
extensionalHCFEquality : {a b : }{h1 h2 : extensionalHCF a b} → (extensionalHCF.c h1 ≡ extensionalHCF.c h2) → h1 ≡ h2
extensionalHCFEquality {a} {b} {record { c = c1 ; c|a = c|a1 ; c|b = c|b1 ; hcfExtension = hcfExtension1 }} {record { c = c2 ; c|a = c|a2 ; c|b = c|b2 ; hcfExtension = hcfExtension2 }} pr rewrite pr = {!!}
-}
record extendedHcf (a b : ) : Set where
field
hcf : hcfData a b
c :
c = hcfData.c hcf
field
extended1 :
extended2 :
extendedProof : (a *N extended1 b *N extended2 +N c) || (a *N extended1 +N c b *N extended2)
divEqualityLemma1 : {a b c : } b zero b *N c +N 0 a a b
divEqualityLemma1 {a} {.0} {c} refl pr = equalityCommutative pr
divEquality : {a b : } a b b a a b
divEquality {a} {b} (divides record { quot = quotAB ; rem = .0 ; pr = prAB ; remIsSmall = _ ; quotSmall = quotSmallAB } refl) (divides record { quot = quot ; rem = .0 ; pr = pr ; remIsSmall = _ ; quotSmall = (inl x) } refl) rewrite additionNIsCommutative (b *N quot) 0 | additionNIsCommutative (a *N quotAB) 0 | equalityCommutative pr | equalityCommutative (multiplicationNIsAssociative b quot quotAB) = res
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} {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
p = lem prAB x
q : quot 1
q = _&&_.fst (productOneImpliesOperandsOne p)
res : b *N quot b
res rewrite q | multiplicationNIsCommutative b 1 | additionNIsCommutative b 0 = refl
divEquality {.0} {.0} (divides record { quot = quotAB ; rem = .0 ; pr = prAB ; remIsSmall = _ ; quotSmall = quotSmallAB } refl) (divides record { quot = quot ; rem = .0 ; pr = refl ; remIsSmall = _ ; quotSmall = (inr (refl ,, snd)) } refl) = refl
hcfWelldefined : {a b : } (ab : hcfData a b) (ab' : hcfData a b) (hcfData.c ab hcfData.c ab')
hcfWelldefined {a} {b} record { c = c ; c|a = c|a ; c|b = c|b ; hcf = hcf } record { c = c' ; c|a = c|a' ; c|b = c|b' ; hcf = hcf' } with hcf c' c|a' c|b'
... | c'DivC with hcf' c c|a c|b
... | cDivC' = divEquality cDivC' c'DivC
reverseHCF : {a b : } (ab : extendedHcf a b) extendedHcf b a
reverseHCF {a} {b} record { hcf = record { c = c ; c|a = c|a ; c|b = c|b ; hcf = hcf } ; extended1 = extended1 ; extended2 = extended2 ; extendedProof = (inl x) } = record { hcf = record { c = c ; c|a = c|b ; c|b = c|a ; hcf = λ x z z₁ hcf x z₁ z } ; extended1 = extended2 ; extended2 = extended1 ; extendedProof = inr (equalityCommutative x) }
reverseHCF {a} {b} record { hcf = record { c = c ; c|a = c|a ; c|b = c|b ; hcf = hcf } ; extended1 = extended1 ; extended2 = extended2 ; extendedProof = (inr x) } = record { hcf = record { c = c ; c|a = c|b ; c|b = c|a ; hcf = λ x z z₁ hcf x z₁ z } ; extended1 = extended2 ; extended2 = extended1 ; extendedProof = inl (equalityCommutative x) }
oneDivN : (a : ) 1 a
oneDivN a = divides (record { quot = a ; rem = zero ; pr = pr ; remIsSmall = inl (succIsPositive zero) ; quotSmall = inl (le zero refl) }) refl
where
pr : (a +N zero) +N zero a
pr rewrite addZeroRight (a +N zero) = addZeroRight a
hcfZero : (a : ) extendedHcf zero a
hcfZero a = record { hcf = record { c = a ; c|a = aDivZero a ; c|b = aDivA a ; hcf = λ _ _ p p } ; extended1 = 0 ; extended2 = 1 ; extendedProof = inr (equalityCommutative (productWithOneRight a))}
hcfOne : (a : ) extendedHcf 1 a
hcfOne a = record { hcf = record { c = 1 ; c|a = aDivA 1 ; c|b = oneDivN a ; hcf = λ _ z _ z } ; extended1 = 1 ; extended2 = 0 ; extendedProof = inl g }
where
g : 1 a *N 0 +N 1
g rewrite multiplicationNIsCommutative a 0 = refl
zeroIsValidRem : (a : ) (0 <N a) || (a 0)
zeroIsValidRem zero = inr refl
zeroIsValidRem (succ a) = inl (succIsPositive a)
dividesBothImpliesDividesSum : {a x y : } a x a y a (x +N y)
dividesBothImpliesDividesSum {a} {x} {y} (divides record { quot = xDivA ; rem = .0 ; pr = prA ; quotSmall = qsm1 } refl) (divides record { quot = quot ; rem = .0 ; pr = pr ; quotSmall = qsm2 } refl) = divides (record { quot = xDivA +N quot ; rem = 0 ; pr = go {a} {x} {y} {xDivA} {quot} pr prA ; remIsSmall = zeroIsValidRem a ; quotSmall = (quotSmall qsm1 qsm2) }) refl
where
go : {a x y quot quot2 : } (a *N quot2 +N zero y) (a *N quot +N zero x) a *N (quot +N quot2) +N zero x +N y
go {a} {x} {y} {quot} {quot2} pr1 pr2 rewrite addZeroRight (a *N quot) = identityOfIndiscernablesLeft (a *N (quot +N quot2)) (x +N y) (a *N (quot +N quot2) +N zero) _≡_ t (equalityCommutative (addZeroRight (a *N (quot +N quot2))))
where
t : a *N (quot +N quot2) x +N y
t rewrite addZeroRight (a *N quot2) = transitivity (productDistributes a quot quot2) p
where
s : a *N quot +N a *N quot2 x +N a *N quot2
s = applyEquality (λ n n +N a *N quot2) pr2
r : x +N a *N quot2 x +N y
r = applyEquality (λ n x +N n) pr1
p : a *N quot +N a *N quot2 x +N y
p = transitivity s r
quotSmall : ((0 <N a) || ((0 a) && (xDivA 0))) ((0 <N a) || ((0 a) && (quot 0))) (0 <N a) || ((0 a) && (xDivA +N quot 0))
quotSmall (inl x1) (inl x2) = inl x1
quotSmall (inl x1) (inr x2) = inl x1
quotSmall (inr x1) (inl x2) = inl x2
quotSmall (inr (a=0 ,, bl)) (inr (_ ,, bl2)) = inr (a=0 ,, ans)
where
ans : xDivA +N quot 0
ans rewrite bl | bl2 = refl
dividesBothImpliesDividesDifference : {a b c : } a b a c (c<b : c <N b) a (subtractionNResult.result (-N (inl c<b)))
dividesBothImpliesDividesDifference {zero} {b} {.0} prab (divides record { quot = quot ; rem = .0 ; pr = refl } refl) c<b = prab
dividesBothImpliesDividesDifference {succ a} {b} {c} (divides record { quot = bDivSA ; rem = .0 ; pr = pr } refl) (divides record { quot = cDivSA ; rem = .0 ; pr = pr2 } refl) c<b rewrite (addZeroRight (succ a *N cDivSA)) | (addZeroRight (succ a *N bDivSA)) = divides (record { quot = subtractionNResult.result bDivSA-cDivSA ; rem = 0 ; pr = identityOfIndiscernablesLeft (succ a *N (subtractionNResult.result bDivSA-cDivSA)) (subtractionNResult.result (-N (inl c<b))) (succ a *N (subtractionNResult.result bDivSA-cDivSA) +N zero) _≡_ (identityOfIndiscernablesLeft (la-ka) (subtractionNResult.result (-N (inl c<b))) (succ a *N (subtractionNResult.result bDivSA-cDivSA)) _≡_ s (equalityCommutative q)) (equalityCommutative (addZeroRight _)) ; remIsSmall = inl (succIsPositive a) ; quotSmall = inl (succIsPositive a) }) refl
where
p : cDivSA <N bDivSA
p rewrite (equalityCommutative pr2) | (equalityCommutative pr) = cancelInequalityLeft {succ a} {cDivSA} {bDivSA} c<b
bDivSA-cDivSA : subtractionNResult cDivSA bDivSA (inl p)
bDivSA-cDivSA = -N {cDivSA} {bDivSA} (inl p)
la-ka = subtractionNResult.result (-N {succ a *N cDivSA} {succ a *N bDivSA} (inl (lessRespectsMultiplicationLeft cDivSA bDivSA (succ a) p (succIsPositive a))))
q : (succ a *N (subtractionNResult.result bDivSA-cDivSA)) la-ka
q = subtractProduct {succ a} {cDivSA} {bDivSA} (succIsPositive a) p
s : la-ka subtractionNResult.result (-N {c} {b} (inl c<b))
s = equivalentSubtraction (succ a *N cDivSA) b (succ a *N bDivSA) c (lessRespectsMultiplicationLeft cDivSA bDivSA (succ a) p (succIsPositive a)) c<b g
where
g : (succ a *N cDivSA) +N b (succ a *N bDivSA) +N c
g rewrite equalityCommutative pr2 | equalityCommutative pr = additionNIsCommutative (cDivSA +N a *N cDivSA) (bDivSA +N a *N bDivSA)
euclidLemma1 : {a b : } (a<b : a <N b) (t : ) a +N b <N t a +N (subtractionNResult.result (-N (inl a<b))) <N t
euclidLemma1 {zero} {b} zero<b t b<t = b<t
euclidLemma1 {succ a} {b} sa<b t sa+b<t = identityOfIndiscernablesLeft (subtractionNResult.result (-N (inl sa<b)) +N succ a) t (succ a +N (subtractionNResult.result (-N (inl sa<b)))) _<N_ q (additionNIsCommutative (subtractionNResult.result (-N (inl sa<b))) (succ a))
where
p : b <N t
p = orderIsTransitive (le a refl) sa+b<t
q : (subtractionNResult.result (-N (inl sa<b))) +N succ a <N t
q = identityOfIndiscernablesLeft b t (subtractionNResult.result (-N (inl sa<b)) +N succ a) _<N_ p (equalityCommutative (addMinus (inl sa<b)))
euclidLemma2 : {a b max : } (succ (a +N b) <N max) b <N max
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))
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))))
euclidLemma4 : (a b c d h : ) (sa<b : (succ a) <N b) (pr : subtractionNResult.result (-N (inl sa<b)) *N c (succ a) *N d +N h) b *N c (succ a) *N (d +N c) +N h
euclidLemma4 a b zero d h sa<b pr rewrite addZeroRight d | productZeroIsZeroRight b | productZeroIsZeroRight (subtractionNResult.result (-N (inl sa<b))) = pr
euclidLemma4 a b (succ c) d h sa<b pr rewrite subtractProduct' (succIsPositive c) sa<b = transitivity q' r'
where
q : (succ c) *N b succ (a +N c *N succ a) +N ((succ a) *N d +N h)
q = moveOneSubtraction {succ (a +N c *N succ a)} {b +N c *N b} {(succ a) *N d +N h} {inl _} pr
q' : b *N succ c succ (a +N c *N succ a) +N ((succ a) *N d +N h)
q' rewrite multiplicationNIsCommutative b (succ c) = q
r' : ((succ c) *N succ a) +N (((succ a) *N d) +N h) ((succ a) *N (d +N succ c)) +N h
r' rewrite equalityCommutative (additionNIsAssociative ((succ c) *N succ a) ((succ a) *N d) h) = applyEquality (λ t t +N h) {((succ c) *N succ a) +N ((succ a) *N d)} {(succ a) *N (d +N succ c)} (go (succ c) (succ a) d)
where
go' : (a b c : ) b *N a +N b *N c b *N (c +N a)
go : (a b c : ) a *N b +N b *N c b *N (c +N a)
go a b c rewrite multiplicationNIsCommutative a b = go' a b c
go' a b c rewrite additionNIsCommutative (b *N a) (b *N c) = equalityCommutative (productDistributes b c a)
euclidLemma5 : (a b c d h : ) (sa<b : (succ a) <N b) (pr : subtractionNResult.result (-N (inl sa<b)) *N c +N h (succ a) *N d) (succ a) *N (d +N c) b *N c +N h
euclidLemma5 a b c d h sa<b pr with (-N (inl sa<b))
euclidLemma5 a b zero d h sa<b pr | record { result = result ; pr = sub } rewrite addZeroRight d | productZeroIsZeroRight b | productZeroIsZeroRight result = equalityCommutative pr
euclidLemma5 a b (succ c) d h sa<b pr | record { result = result ; pr = sub } rewrite subtractProduct' (succIsPositive c) sa<b | equalityCommutative sub = pv''
where
p : succ a *N d result *N succ c +N h
p = equalityCommutative pr
p' : a *N succ c +N succ a *N d (a *N succ c) +N ((result *N succ c) +N h)
p' = applyEquality (λ t a *N succ c +N t) p
p'' : a *N succ c +N succ a *N d (a *N succ c +N result *N succ c) +N h
p'' rewrite (additionNIsAssociative (a *N succ c) (result *N succ c) h) = p'
p''' : a *N succ c +N succ a *N d (a +N result) *N succ c +N h
p''' rewrite productDistributes' a result (succ c) = p''
pv : c +N (a *N succ c +N succ a *N d) (c +N (a +N result) *N succ c) +N h
pv rewrite additionNIsAssociative c ((a +N result) *N succ c) h = applyEquality (λ t c +N t) p'''
pv' : (succ c) +N (a *N succ c +N succ a *N d) succ ((c +N (a +N result) *N succ c) +N h)
pv' = applyEquality succ pv
pv'' : (succ a) *N (d +N succ c) succ ((c +N (a +N result) *N succ c) +N h)
pv'' = identityOfIndiscernablesLeft ((succ c) +N (a *N succ c +N succ a *N d)) _ ((succ a) *N (d +N succ c)) _≡_ pv' (go a c d)
where
go : (a c d : ) (succ c) +N (a *N succ c +N ((succ a) *N d)) (succ a) *N (d +N succ c)
go a c d rewrite equalityCommutative (additionNIsAssociative (succ c) (a *N succ c) ((succ a) *N d)) = go'
where
go' : (succ a) *N (succ c) +N (succ a) *N d (succ a) *N (d +N succ c)
go' rewrite additionNIsCommutative d (succ c) = equalityCommutative (productDistributes (succ a) (succ c) d)
euclid : (a b : ) extendedHcf a b
euclid a b = inducted (succ a +N b) a b (a<SuccA (a +N b))
where
P : Set
P sum = (a b : ) a +N b <N sum extendedHcf a b
go'' : {a b : } (maxsum : ) (a <N b) (a +N b <N maxsum) ( y y <N maxsum P y) extendedHcf a b
go'' {zero} {b} maxSum zero<b b<maxsum indHyp = hcfZero b
go'' {1} {b} maxSum 1<b b<maxsum indHyp = hcfOne b
go'' {succ (succ a)} {b} maxSum ssa<b ssa+b<maxsum indHyp with (indHyp (succ b) (euclidLemma3 {a} {b} {maxSum} ssa+b<maxsum)) (subtractionNResult.result (-N (inl ssa<b))) (succ (succ a)) (identityOfIndiscernablesLeft b (succ b) (subtractionNResult.result (-N (inl ssa<b)) +N succ (succ a)) _<N_ (a<SuccA b) (equalityCommutative (addMinus (inl ssa<b))))
go'' {succ (succ a)} {b} maxSum ssa<b ssa+b<maxsum indHyp | record { hcf = record { c = c ; c|a = c|a ; c|b = c|b ; hcf = hcf } ; extended1 = extended1 ; extended2 = extended2 ; extendedProof = inl extendedProof } = record { hcf = record { c = c ; c|a = c|b ; c|b = hcfDivB'' ; hcf = λ div prDivSSA prDivB hcf div (dividesBothImpliesDividesDifference prDivB prDivSSA ssa<b) prDivSSA } ; extended2 = extended1; extended1 = extended2 +N extended1 ; extendedProof = inr (equalityCommutative (euclidLemma4 (succ a) b extended1 extended2 c ssa<b extendedProof)) }
where
hcfDivB : c ((succ (succ a)) +N (subtractionNResult.result (-N (inl ssa<b))))
hcfDivB = dividesBothImpliesDividesSum {c} {succ (succ a)} { subtractionNResult.result (-N (inl ssa<b))} c|b c|a
hcfDivB' : c ((subtractionNResult.result (-N (inl ssa<b))) +N (succ (succ a)))
hcfDivB' = identityOfIndiscernablesRight c ((succ (succ a)) +N (subtractionNResult.result (-N (inl ssa<b)))) ((subtractionNResult.result (-N (inl ssa<b))) +N (succ (succ a))) __ hcfDivB (additionNIsCommutative (succ (succ a)) ( subtractionNResult.result (-N (inl ssa<b))))
hcfDivB'' : c b
hcfDivB'' = identityOfIndiscernablesRight c ((subtractionNResult.result (-N (inl ssa<b))) +N (succ (succ a))) b __ hcfDivB' (addMinus (inl ssa<b))
go'' {succ (succ a)} {b} maxSum ssa<b ssa+b<maxsum indHyp | record { hcf = record { c = c ; c|a = c|a ; c|b = c|b ; hcf = hcf } ; extended1 = extended1 ; extended2 = extended2 ; extendedProof = inr extendedProof } = record { hcf = record { c = c ; c|a = c|b ; c|b = hcfDivB'' ; hcf = λ div prDivSSA prDivB hcf div (dividesBothImpliesDividesDifference prDivB prDivSSA ssa<b) prDivSSA } ; extended2 = extended1; extended1 = extended2 +N extended1 ; extendedProof = inl (euclidLemma5 (succ a) b extended1 extended2 c ssa<b extendedProof) }
where
hcfDivB : c ((succ (succ a)) +N (subtractionNResult.result (-N (inl ssa<b))))
hcfDivB = dividesBothImpliesDividesSum {c} {succ (succ a)} { subtractionNResult.result (-N (inl ssa<b))} c|b c|a
hcfDivB' : c ((subtractionNResult.result (-N (inl ssa<b))) +N (succ (succ a)))
hcfDivB' = identityOfIndiscernablesRight c ((succ (succ a)) +N (subtractionNResult.result (-N (inl ssa<b)))) ((subtractionNResult.result (-N (inl ssa<b))) +N (succ (succ a))) __ hcfDivB (additionNIsCommutative (succ (succ a)) ( subtractionNResult.result (-N (inl ssa<b))))
hcfDivB'' : c b
hcfDivB'' = identityOfIndiscernablesRight c ((subtractionNResult.result (-N (inl ssa<b))) +N (succ (succ a))) b __ hcfDivB' (addMinus (inl ssa<b))
go' : (maxSum a b : ) (a +N b <N maxSum) ( y y <N maxSum P y) extendedHcf a b
go' maxSum a b a+b<maxsum indHyp with orderIsTotal a b
go' maxSum a b a+b<maxsum indHyp | inl (inl a<b) = go'' maxSum a<b a+b<maxsum indHyp
go' maxSum a b a+b<maxsum indHyp | inl (inr b<a) = reverseHCF (go'' maxSum b<a (identityOfIndiscernablesLeft (a +N b) maxSum (b +N a) _<N_ a+b<maxsum (additionNIsCommutative a b)) indHyp)
go' maxSum a .a _ indHyp | inr refl = record { hcf = record { c = a ; c|a = aDivA a ; c|b = aDivA a ; hcf = λ _ _ z z } ; extended1 = 0 ; extended2 = 1 ; extendedProof = inr s}
where
s : a *N zero +N a a *N 1
s rewrite multiplicationNIsCommutative a zero | productWithOneRight a = refl
go : x ( y y <N x P y) P x
go maxSum indHyp = λ a b a+b<maxSum go' maxSum a b a+b<maxSum indHyp
inducted : x P x
inducted = rec <NWellfounded P go
divisorIsSmaller : {a b : } a succ b succ b <N a False
divisorIsSmaller {a} {b} (divides record { quot = zero ; rem = .0 ; pr = pr } refl) sb<a rewrite addZeroRight (a *N zero) = go
where
go : False
go rewrite productZeroIsZeroRight a = naughtE pr
divisorIsSmaller {a} {b} (divides record { quot = (succ quot) ; rem = .0 ; pr = pr } refl) sb<a rewrite addZeroRight (a *N succ quot) = go
where
go : False
go rewrite equalityCommutative pr = go'
where
go' : False
go' rewrite multiplicationNIsCommutative a (succ quot) = cannotAddAndEnlarge' sb<a
primeDivisorIs1OrP : {a p : } (prime : Prime p) (a p) (a 1) || (a p)
primeDivisorIs1OrP {zero} {zero} prime a|p = inr refl
primeDivisorIs1OrP {zero} {succ p} prime a|p = exFalso (zeroDividesNothing p a|p)
primeDivisorIs1OrP {succ zero} {p} prime a|p = inl refl
primeDivisorIs1OrP {succ (succ a)} {p} prime a|p with orderIsTotal (succ (succ a)) p
primeDivisorIs1OrP {succ (succ a)} {p} prime a|p | inl (inl ssa<p) = go p prime a|p ssa<p
where
go : (n : ) Prime n succ (succ a) n succ (succ a) <N n (succ (succ a) 1) || (succ (succ a) p)
go zero pr x|n n<n = exFalso (zeroIsNotPrime pr)
go (succ zero) pr x|n n<n = exFalso (oneIsNotPrime pr)
go (succ (succ n)) pr x|n n<n = inl ((Prime.pr pr) {succ (succ a)} x|n n<n (succIsPositive (succ a)))
primeDivisorIs1OrP {succ (succ a)} {zero} prime a|p | inl (inr x) = exFalso (zeroIsNotPrime prime)
primeDivisorIs1OrP {succ (succ a)} {succ p} prime a|p | inl (inr x) = exFalso (divisorIsSmaller {succ (succ a)} {p} a|p x)
primeDivisorIs1OrP {succ (succ a)} {p} prime a|p | inr x = inr x
hcfPrimeIsOne' : {p : } {a : } (Prime p) (0 <N divisionAlgResult.rem (divisionAlg p a)) (extendedHcf.c (euclid a p) 1) || (extendedHcf.c (euclid a p) p)
hcfPrimeIsOne' {p} {a} pPrime pCoprimeA with euclid a p
hcfPrimeIsOne' {p} {a} pPrime pCoprimeA | record { hcf = record { c = hcf ; c|a = hcf|a ; c|b = hcf|p ; hcf = hcfPr } } with divisionAlg p a
hcfPrimeIsOne' {p} {a} pPrime pCoprimeA | record { hcf = record { c = hcf ; c|a = hcf|a ; c|b = hcf|p ; hcf = hcfPr } } | record { quot = quot ; rem = rem ; pr = prPDivA } with primeDivisorIs1OrP pPrime hcf|p
hcfPrimeIsOne' {p} {a} pPrime pCoprimeA | record { hcf = record { c = hcf ; c|a = hcf|a ; c|b = hcf|p ; hcf = hcfPr } } | record { quot = quot ; rem = rem ; pr = prPDivA } | inl x = inl x
hcfPrimeIsOne' {p} {a} pPrime pCoprimeA | record { hcf = record { c = hcf ; c|a = hcf|a ; c|b = hcf|p ; hcf = hcfPr } } | record { quot = quot ; rem = rem ; pr = prPDivA } | inr x = inr x
divisionDecidable : (a b : ) (a b) || ((a b) False)
divisionDecidable zero zero = inl (aDivA zero)
divisionDecidable zero (succ b) = inr f
where
f : zero succ b False
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 = inl p } = inr f
where
f : (succ a) b False
f (divides record { quot = b/a' ; rem = .0 ; pr = pr } refl) rewrite addZeroRight ((succ a) *N b/a') = naughtE (modUniqueLemma {zero} {succ rem} {succ a} b/a' b/a (succIsPositive a) p comp')
where
comp : (succ a) *N b/a' (succ a) *N b/a +N succ rem
comp = transitivity pr (equalityCommutative prANotDivB)
comp' : (succ a) *N b/a' +N zero (succ a) *N b/a +N succ rem
comp' rewrite addZeroRight (succ a *N b/a') = comp
doesNotDivideImpliesNonzeroRem : (a b : ) ((a b) False) 0 <N divisionAlgResult.rem (divisionAlg a b)
doesNotDivideImpliesNonzeroRem a b pr with divisionAlg a b
doesNotDivideImpliesNonzeroRem a b pr | record { quot = quot ; rem = rem ; pr = divAlgPr ; remIsSmall = remIsSmall } with zeroIsValidRem rem
doesNotDivideImpliesNonzeroRem a b pr | record { quot = quot ; rem = rem ; pr = divAlgPr ; remIsSmall = remIsSmall } | inl x = x
doesNotDivideImpliesNonzeroRem a b pr | record { quot = quot ; rem = rem ; pr = divAlgPr ; remIsSmall = remIsSmall ; quotSmall = quotSmall } | inr x = exFalso (pr aDivB)
where
aDivB : a b
aDivB = divides (record { quot = quot ; rem = rem ; pr = divAlgPr ; remIsSmall = remIsSmall ; quotSmall = quotSmall }) x
hcfPrimeIsOne : {p : } {a : } (Prime p) ((p a) False) extendedHcf.c (euclid a p) 1
hcfPrimeIsOne {p} {a} pPrime pr with hcfPrimeIsOne' {p} {a} pPrime (doesNotDivideImpliesNonzeroRem p a pr)
hcfPrimeIsOne {p} {a} pPrime pr | inl x = x
hcfPrimeIsOne {p} {a} pPrime pr | inr x with euclid a p
hcfPrimeIsOne {p} {a} pPrime pr | inr x | record { hcf = record { c = c ; c|a = c|a ; c|b = c|b ; hcf = hcf } ; extended1 = extended1 ; extended2 = extended2 ; extendedProof = extendedProof } rewrite x = exFalso (pr c|a)
reduceEquationMod : {a b c : } (d : ) (a b) (a c) b c +N d a d
reduceEquationMod {a} {b} {c} 0 a|b a|c pr = aDivZero a
reduceEquationMod {a} {b} {c} (succ d) (divides record { quot = b/a ; rem = .0 ; pr = prb/a ; remIsSmall = r1 ; quotSmall = qSm1 } refl) (divides record { quot = c/a ; rem = .0 ; pr = prc/a ; remIsSmall = r2 ; quotSmall = qSm2 } refl) b=c+d = identityOfIndiscernablesRight a (subtractionNResult.result (-N (inl c<b))) (succ d) __ a|b-c ex'
where
c<b : c <N b
c<b rewrite succExtracts c d | additionNIsCommutative c d = le d (equalityCommutative b=c+d)
a|b-c : a subtractionNResult.result (-N (inl c<b))
a|b-c = dividesBothImpliesDividesDifference (divides record { quot = b/a ; rem = 0 ; pr = prb/a ; remIsSmall = r1 ; quotSmall = qSm1 } refl) (divides record { quot = c/a ; rem = 0 ; pr = prc/a ; remIsSmall = r2 ; quotSmall = qSm2 } refl) c<b
ex : subtractionNResult.result (-N {c} {b} (inl c<b)) subtractionNResult.result (-N {0} {succ d} (inl (succIsPositive d)))
ex = equivalentSubtraction c (succ d) b 0 (c<b) (succIsPositive d) (equalityCommutative (identityOfIndiscernablesLeft b (c +N succ d) (b +N 0) _≡_ b=c+d (equalityCommutative (addZeroRight b))))
ex' : subtractionNResult.result (-N {c} {b} (inl c<b)) succ d
ex' = identityOfIndiscernablesRight (subtractionNResult.result (-N (inl c<b))) (subtractionNResult.result (-N (inl (succIsPositive d)))) (succ d) _≡_ ex refl
primesArePrime : {p : } {a b : } (Prime p) p (a *N b) (p a) || (p b)
primesArePrime {p} {a} {b} pPrime pr with divisionDecidable p a
primesArePrime {p} {a} {b} pPrime pr | inl p|a = inl p|a
primesArePrime {p} {a} {b} pPrime (divides record {quot = ab/p ; rem = .0 ; pr = p|ab ; remIsSmall = _ ; quotSmall = quotSmall } refl) | inr notp|a = inr (answer ex'')
where
euc : extendedHcf a p
euc = euclid a p
h : extendedHcf.c euc 1
h = hcfPrimeIsOne {p} {a} pPrime notp|a
x = extendedHcf.extended1 euc
y = extendedHcf.extended2 euc
extended : ((a *N x) p *N y +N extendedHcf.c euc) || (a *N x +N extendedHcf.c euc p *N y)
extended = extendedHcf.extendedProof euc
extended' : (a *N x p *N y +N 1) || (a *N x +N 1 p *N y)
extended' rewrite equalityCommutative h = extended
extended'' : ((a *N x p *N y +N 1) || (a *N x +N 1 p *N y)) (b *N (a *N x) b *N (p *N y +N 1)) || (b *N (a *N x +N 1) b *N (p *N y))
extended'' (inl z) = inl (applyEquality (λ t b *N t) z)
extended'' (inr z) = inr (applyEquality (λ t b *N t) z)
ex : (b *N (a *N x) b *N (p *N y +N 1)) || (b *N (a *N x +N 1) b *N (p *N y)) ((b *N a) *N x (b *N (p *N y) +N b)) || (((b *N a) *N x +N b) b *N (p *N y))
ex (inl z) rewrite multiplicationNIsAssociative b a x | productDistributes b (p *N y) 1 | productWithOneRight b = inl z
ex (inr z) rewrite productDistributes b (a *N x) 1 | multiplicationNIsAssociative b a x | productWithOneRight b = inr z
ex' : ((a *N b) *N x ((p *N y) *N b +N b)) || (((a *N b) *N x +N b) (p *N y) *N b)
ex' rewrite multiplicationNIsCommutative a b | multiplicationNIsCommutative (p *N y) b = ex (extended'' extended')
ex'' : ((a *N b) *N x (p *N (y *N b) +N b)) || (((a *N b) *N x +N b) p *N (y *N b))
ex'' rewrite multiplicationNIsAssociative (p) y b = ex'
inter1 : p (a *N b) *N x
inter1 = divides (record {quot = ab/p *N x ; rem = 0 ; pr = g ; remIsSmall = zeroIsValidRem p ; quotSmall = qsm quotSmall}) refl
where
g' : p *N ab/p a *N b
g' rewrite addZeroRight (p *N ab/p) = p|ab
g'' : p *N (ab/p *N x) (a *N b) *N x
g'' rewrite multiplicationNIsAssociative (p) ab/p x = applyEquality (λ t t *N x) g'
g : p *N (ab/p *N x) +N 0 (a *N b) *N x
g rewrite addZeroRight (p *N (ab/p *N x)) = g''
qsm : ((0 <N p) || ((0 p) && (ab/p 0))) (0 <N p) || ((0 p) && (ab/p *N x 0))
qsm (inl pr) = inl pr
qsm (inr (pr1 ,, pr2)) = inr (pr1 ,, blah)
where
blah : ab/p *N x 0
blah rewrite pr2 = refl
inter2 : ((0 <N p) || ((0 p) && (ab/p 0))) p (p *N (y *N b))
inter2 (inl 0<p) = divides (record {quot = y *N b ; rem = 0 ; pr = addZeroRight (p *N (y *N b)) ; remIsSmall = zeroIsValidRem p ; quotSmall = inl 0<p }) refl
inter2 (inr (0=p ,, ab/p=0)) = divides (record {quot = y *N b ; rem = 0 ; pr = addZeroRight (p *N (y *N b)) ; remIsSmall = zeroIsValidRem p ; quotSmall = inr (0=p ,, blah) }) refl
where
oneZero : (a 0) || (b 0)
oneZero rewrite additionNIsCommutative (p *N ab/p) 0 | ab/p=0 | equalityCommutative 0=p = productZeroImpliesOperandZero (equalityCommutative p|ab)
blah : y *N b 0
blah with oneZero
blah | inl x1 = exFalso (notp|a p|a)
where
p|a : p a
p|a rewrite x1 = aDivZero p
blah | inr x1 rewrite x1 = productZeroIsZeroRight y
answer : ((a *N b) *N x (p *N (y *N b) +N b)) || (((a *N b) *N x +N b) p *N (y *N b)) (p b)
answer (inl z) = reduceEquationMod {p} b inter1 (inter2 quotSmall) z
answer (inr z) = reduceEquationMod {p} b (inter2 quotSmall) inter1 (equalityCommutative z)
primesAreBiggerThanOne : {p : } Prime p (1 <N p)
primesAreBiggerThanOne {zero} record { p>1 = (le x ()) ; pr = pr }
primesAreBiggerThanOne {succ zero} pr = exFalso (oneIsNotPrime pr)
primesAreBiggerThanOne {succ (succ p)} pr = succPreservesInequality (succIsPositive p)
primesAreBiggerThanZero : {p : } Prime p 0 <N p
primesAreBiggerThanZero {p} pr = orderIsTransitive (succIsPositive 0) (primesAreBiggerThanOne pr)
record notDividedByLessThan (a : ) (firstPossibleDivisor : ) : Set where
field
previousDoNotDivide : x 1 <N x x <N firstPossibleDivisor x a False
alternativePrime : {a : } 1 <N a notDividedByLessThan a a Prime a
alternativePrime {a} 1<a record { previousDoNotDivide = previousDoNotDivide } = record { pr = pr ; p>1 = 1<a}
where
pr : {x : } (x|a : x a) (x<a : x <N a) (0<x : zero <N x) x 1
pr {zero} _ _ (le x ())
pr {succ zero} _ _ _ = refl
pr {succ (succ x)} x|a x<a 0<x = exFalso (previousDoNotDivide (succ (succ x)) (succPreservesInequality (succIsPositive x)) x<a x|a)
divisibilityTransitive : {a b c : } a b b c a c
divisibilityTransitive {a} {b} {c} (divides record { quot = b/a ; rem = .0 ; pr = prDivAB ; remIsSmall = remIsSmallAB ; quotSmall = quotSmall } refl) (divides record { quot = c/b ; rem = .0 ; pr = prDivBC ; remIsSmall = remIsSmallBC } refl) = divides record { quot = b/a *N c/b ; rem = 0 ; pr = p ; remIsSmall = zeroIsValidRem a ; quotSmall = qsm quotSmall } refl
where
p : a *N (b/a *N c/b) +N 0 c
p rewrite addZeroRight (a *N (b/a *N c/b)) | addZeroRight (b *N c/b) | addZeroRight (a *N b/a) | multiplicationNIsAssociative a b/a c/b | prDivAB | prDivBC = refl
qsm : ((0 <N a) || (0 a) && (b/a 0)) ((0 <N a) || (0 a) && (b/a *N c/b 0))
qsm (inl x) = inl x
qsm (inr (p1 ,, p2)) = inr (p1 ,, blah)
where
blah : b/a *N c/b 0
blah rewrite p2 = refl
compositeOrPrimeLemma : {a b : } notDividedByLessThan b a a b {i : } (i a) (i <N a) (0 <N i) i 1
compositeOrPrimeLemma {a} {b} record { previousDoNotDivide = previousDoNotDivide } a|b {zero} i|a i<a 0<i = exFalso (lessIrreflexive 0<i)
compositeOrPrimeLemma {a} {b} record { previousDoNotDivide = previousDoNotDivide } a|b {succ zero} i|a i<a 0<i = refl
compositeOrPrimeLemma {a} {b} record { previousDoNotDivide = previousDoNotDivide } a|b {succ (succ i)} i|a i<a 0<i = exFalso (previousDoNotDivide (succ (succ i)) (succPreservesInequality (succIsPositive i)) i<a (divisibilityTransitive i|a a|b) )
compositeOrPrime : (a : ) (1 <N a) (Composite a) || (Prime a)
compositeOrPrime a pr = go''' go''
where
base : notDividedByLessThan a 2
base = record { previousDoNotDivide = λ x 1<x x<2 _ noIntegersBetweenXAndSuccX 1 1<x x<2 }
go : {firstPoss : } notDividedByLessThan a firstPoss ((notDividedByLessThan a (succ firstPoss)) || ((firstPoss a) && (firstPoss <N a))) || (firstPoss a)
go' : (firstPoss : ) (((notDividedByLessThan a firstPoss) || (Composite a))) || (notDividedByLessThan a a)
go'' : (notDividedByLessThan a a) || (Composite a)
go'' with go' a
... | inr x = inl x
... | inl x = x
go''' : ((notDividedByLessThan a a) || (Composite a)) ((Composite a) || (Prime a))
go''' (inl x) = inr (alternativePrime pr x)
go''' (inr x) = inl x
go' (zero) = inl (inl (record { previousDoNotDivide = λ x 1<x x<0 _ zeroNeverGreater x<0 }))
go' (succ 0) = inl (inl (record { previousDoNotDivide = λ x 1<x x<1 _ orderIsIrreflexive x<1 1<x }))
go' (succ (succ zero)) = inl (inl base)
go' (succ (succ (succ firstPoss))) with go' (succ (succ firstPoss))
go' (succ (succ (succ firstPoss))) | inl (inl x) with go {succ (succ firstPoss)} x
go' (succ (succ (succ firstPoss))) | inl (inl x) | inl (inl x1) = inl (inl x1)
go' (succ (succ (succ firstPoss))) | inl (inl x) | inl (inr x1) = inl (inr record { noSmallerDivisors = λ i i<ssFP 1<i i|a notDividedByLessThan.previousDoNotDivide x i 1<i i<ssFP i|a ; n>1 = pr ; divisor = succ (succ firstPoss) ; dividesN = _&&_.fst x1 ; divisorLessN = _&&_.snd x1 ; divisorNot1 = succPreservesInequality (succIsPositive firstPoss) ; divisorPrime = record { p>1 = succPreservesInequality (succIsPositive firstPoss) ; pr = compositeOrPrimeLemma {succ (succ firstPoss)} {a} x (_&&_.fst x1) } })
go' (succ (succ (succ firstPoss))) | inl (inl x) | inr y rewrite y = inr x
go' (succ (succ (succ firstPoss))) | inl (inr x) = inl (inr x)
go' (succ (succ (succ firstPoss))) | inr x = inr x
go {zero} pr = inl (inl (record { previousDoNotDivide = λ x 1<x x<1 _ orderIsIrreflexive x<1 1<x}))
go {succ firstPoss} knownCoprime with orderIsTotal (succ firstPoss) a
go {succ firstPoss} knownCoprime | inr x = inr x
go {succ firstPoss} knownCoprime | inl (inl sFP<a) with divisionAlg (succ firstPoss) a
go {succ firstPoss} knownCoprime | inl (inl sFP<a) | record { quot = quot ; rem = zero ; pr = pr ; remIsSmall = remIsSmall } = inl (inr record { fst = (divides (record { quot = quot ; rem = zero ; remIsSmall = remIsSmall ; pr = pr ; quotSmall = inl (succIsPositive firstPoss) }) refl) ; snd = sFP<a })
go {succ firstPoss} knownCoprime | inl (inl sFP<a) | record { quot = quot ; rem = succ rem ; pr = pr ; remIsSmall = remIsSmall } = inl next
where
previous : x 1 <N x x <N succ firstPoss x a False
previous = notDividedByLessThan.previousDoNotDivide knownCoprime
next : notDividedByLessThan a (succ (succ firstPoss)) || (((succ firstPoss) a) && (succ firstPoss <N a))
next with divisionAlg (succ firstPoss) a
next | record { quot = quot ; rem = zero ; pr = pr ; remIsSmall = remIsSmall } = inr (record { fst = divides record { quot = quot ; rem = zero ; pr = pr ; remIsSmall = remIsSmall ; quotSmall = inl (succIsPositive firstPoss) } refl ; snd = sFP<a } )
next | record { quot = quot ; rem = succ rem ; pr = pr ; remIsSmall = remIsSmall } = inl record { previousDoNotDivide = (next' record { quot = quot ; rem = succ rem ; pr = pr ; remIsSmall = remIsSmall ; quotSmall = inl (succIsPositive firstPoss) } (succIsPositive rem)) }
where
next' : (res : divisionAlgResult (succ firstPoss) a) (pr : 0 <N divisionAlgResult.rem res) (x : ) 1 <N x x <N succ (succ firstPoss) x a False
next' (res) (prDiv) x 1<x x<ssFirstposs x|a with orderIsTotal x (succ firstPoss)
next' (res) (prDiv) x 1<x x<ssFirstposs x|a | inl (inl x<sFirstPoss) = previous x 1<x x<sFirstPoss x|a
next' (res) (prDiv) x 1<x x<ssFirstposs x|a | inl (inr sFirstPoss<x) = noIntegersBetweenXAndSuccX (succ firstPoss) sFirstPoss<x x<ssFirstposs
next' res prDiv x 1<x x<ssFirstposs (divides res1 x1) | inr x=sFirstPoss rewrite equalityCommutative x=sFirstPoss = g
where
g : False
g with modIsUnique res res1
... | r rewrite r = lessImpliesNotEqual prDiv (equalityCommutative x1)
go {succ firstPoss} record { previousDoNotDivide = previousDoNotDivide } | inl (inr a<sFP) = exFalso (previousDoNotDivide a pr a<sFP (aDivA a))
primeDivPrimeImpliesEqual : {p1 p2 : } Prime p1 Prime p2 p1 p2 p1 p2
primeDivPrimeImpliesEqual {p1} {p2} pr1 pr2 p1|p2 with orderIsTotal p1 p2
primeDivPrimeImpliesEqual {p1} {p2} pr1 record { p>1 = p>1 ; pr = pr } p1|p2 | inl (inl p1<p2) with pr p1|p2 p1<p2 (primesAreBiggerThanZero {p1} pr1)
... | p1=1 = exFalso (oneIsNotPrime contr)
where
contr : Prime 1
contr rewrite p1=1 = pr1
primeDivPrimeImpliesEqual {p1} {zero} pr1 pr2 p1|p2 | inl (inr p1>p2) = exFalso (zeroIsNotPrime pr2)
primeDivPrimeImpliesEqual {p1} {succ p2} pr1 pr2 p1|p2 | inl (inr p1>p2) = exFalso (divisorIsSmaller p1|p2 p1>p2)
primeDivPrimeImpliesEqual {p1} {p2} pr1 pr2 p1|p2 | inr p1=p2 = p1=p2
mult1Lemma : {a b : } a *N succ b 1 (a 1) && (b 0)
mult1Lemma {a} {b} pr = record { fst = _&&_.fst p ; snd = q}
where
p : (a 1) && (succ b 1)
p = productOneImpliesOperandsOne pr
q : b zero
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 = (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
notSmallerMeansGE {a} {b} notA<b with orderIsTotal a b
notSmallerMeansGE {a} {b} notA<b | inl (inl x) = exFalso (notA<b x)
notSmallerMeansGE {a} {b} notA<b | inl (inr x) = inl x
notSmallerMeansGE {a} {b} notA<b | inr x = inr (equalityCommutative x)

194
Rationals.agda Normal file
View File

@@ -0,0 +1,194 @@
{-# OPTIONS --safe --warning=error #-}
open import LogicalFormulae
open import Naturals
open import Integers
open import Groups
open import Rings
open import Fields
open import PrimeNumbers
open import Setoids
open import Functions
module Rationals where
data Sign : Set where
Positive : Sign
Negative : Sign
record : Set where
field
consNum :
consDenom :
0<consDenom : 0 <N consDenom
sign : Sign
sign with consNum
sign | nonneg x = Positive
sign | negSucc x = Negative
numerator :
numerator = {!!}
denominator :
denominator = {!!}
denomPos : 0 <N denominator
denomPos = {!!}
coprime : Coprime numerator denominator
coprime = {!!}
multZero : (c : ) (nonneg 0 *Z c nonneg 0)
multZero c with convertZ c
... | bl = refl
zeroMultLemma : {b c : } {d : } (0 <N d) (nonneg zero *Z c) b *Z nonneg d b nonneg 0
zeroMultLemma {nonneg b} {nonneg c} {zero} (le x ()) pr
zeroMultLemma {nonneg zero} {nonneg c} {succ d} _ pr = refl
zeroMultLemma {nonneg (succ b)} {nonneg c} {succ d} _ pr rewrite multZero (nonneg c) = naughtE (nonnegInjective pr)
zeroMultLemma {nonneg zero} {negSucc c} {d} 0<d pr = refl
zeroMultLemma {nonneg (succ b)} {negSucc c} {zero} (le x ()) pr
zeroMultLemma {nonneg (succ b)} {negSucc c} {succ d} _ pr rewrite multZero (negSucc c) = naughtE (nonnegInjective pr)
zeroMultLemma {negSucc b} {c} {zero} ()
zeroMultLemma {negSucc b} {c} {succ d} _ pr rewrite multZero c = exFalso (done pr)
where
done : {a b : } nonneg a negSucc b False
done ()
cancelIntegerMultNegsucc : {a b : } {c : } negSucc c *Z a negSucc c *Z b a b
cancelIntegerMultNegsucc {nonneg a} {nonneg b} {c} pr = {!!}
cancelIntegerMultNegsucc {nonneg a} {negSucc b} {c} pr = {!!}
cancelIntegerMultNegsucc {negSucc a} {b} {c} pr = {!!}
Setoid : Setoid
(Setoid Setoid. record { consNum = numA ; consDenom = denomA ; 0<consDenom = 0<denomA }) record { consNum = numB ; consDenom = denomB ; 0<consDenom = 0<denomB } = numA *Z (nonneg denomB) numB *Z (nonneg denomA)
Reflexive.reflexive (Equivalence.reflexive (Setoid.eq Setoid)) = refl
Symmetric.symmetric (Equivalence.symmetric (Setoid.eq Setoid)) {b} {c} pr = equalityCommutative pr
Transitive.transitive (Equivalence.transitive (Setoid.eq Setoid)) {record { consNum = a ; consDenom = b ; 0<consDenom = 0<b }} {record { consNum = nonneg zero ; consDenom = d ; 0<consDenom = 0<d }} {record { consNum = e ; consDenom = f ; 0<consDenom = 0<f }} pr1 pr2 = transitivity af=0 (equalityCommutative eb=0)
where
e=0 : e nonneg 0
e=0 = zeroMultLemma {e} {nonneg f} 0<d pr2
a=0 : a nonneg 0
a=0 = zeroMultLemma {a} {nonneg b} 0<d (equalityCommutative pr1)
af=0 : a *Z nonneg f nonneg 0
af=0 rewrite a=0 = refl
eb=0 : e *Z nonneg b nonneg 0
eb=0 rewrite e=0 = refl
Transitive.transitive (Equivalence.transitive (Setoid.eq Setoid)) {record { consNum = a ; consDenom = b ; 0<consDenom = 0<b }} {record { consNum = nonneg (succ x) ; consDenom = d ; 0<consDenom = 0<d }} {record { consNum = e ; consDenom = f ; 0<consDenom = 0<f }} pr1 pr2 = {!!}
Transitive.transitive (Equivalence.transitive (Setoid.eq Setoid)) {record { consNum = a ; consDenom = b ; 0<consDenom = 0<b }} {record { consNum = negSucc x ; consDenom = d ; 0<consDenom = 0<d }} {record { consNum = e ; consDenom = f ; 0<consDenom = 0<f }} pr1 pr2 = {!!}
_+Q_ :
record { consNum = numA ; consDenom = denomA ; 0<consDenom = prA } +Q record { consNum = numB ; consDenom = denomB ; 0<consDenom = prB } = record { consNum = numA +Z numB ; consDenom = denomA *N denomB ; 0<consDenom = productNonzeroIsNonzero prA prB }
0Q :
0Q = record { consNum = nonneg 0 ; consDenom = 1 ; 0<consDenom = le zero refl }
negateQ :
.consNum (negateQ record { consNum = consNum ; consDenom = consDenom ; 0<consDenom = 0<consDenom }) = Group.inverse (Ring.additiveGroup zRing) consNum
.consDenom (negateQ record { consNum = consNum ; consDenom = consDenom ; 0<consDenom = 0<consDenom }) = consDenom
.0<consDenom (negateQ record { consNum = consNum ; consDenom = consDenom ; 0<consDenom = 0<consDenom }) = 0<consDenom
qRightInverse : (a : ) (a +Q 0Q) a
qRightInverse record { consNum = consNum ; consDenom = consDenom ; 0<consDenom = 0<consDenom } = {!!}
qGroup : Group {_} {}
Group.setoid qGroup = Setoid
Group.wellDefined qGroup pr1 pr2 = {!!}
Group._·_ qGroup = _+Q_
Group.identity qGroup = 0Q
Group.inverse qGroup = negateQ
Group.multAssoc qGroup = {!!}
Group.multIdentRight qGroup = {!!}
Group.multIdentLeft qGroup = {!!}
Group.invLeft qGroup = {!!}
Group.invRight qGroup = {!!}
{-
data Sign : Set where
Positive : Sign
Negative : Sign
record Abs : Set where
field
numerator :
denominator :
denomPos : 0 <N denominator
hcf : hcfData denominator numerator
coprime : 0 <N hcfData.c hcf
data : Set where
0Q :
positiveQ : Abs
negativeQ : Abs
equalityAbsQ : (a : Abs) → (b : Abs) → (numsEq : Abs.numerator a ≡ Abs.numerator b) → (denomsEq : Abs.denominator a ≡ Abs.denominator b) → a ≡ b
equalityAbsQ record { numerator = numeratorA ; denominator = denominatorA ; denomPos = denomPosA ; hcf = hcfA ; coprime = coprimeA } record { numerator = .numeratorA ; denominator = .denominatorA ; denomPos = denomPos ; hcf = hcf ; coprime = coprime } refl refl rewrite <NRefl denomPosA denomPos = {!!}
inject : (a : ) →
inject (nonneg zero) = 0Q
inject (nonneg (succ x)) = positiveQ record { numerator = x ; denominator = 1 ; denomPos = le zero refl ; hcf = record { c = 1 ; c|a = aDivA 1 ; c|b = oneDivN x ; hcf = λ i i|1 i|x → i|1 } ; coprime = le zero refl }
where
q : 1 ≡ x *N 0 +N 1
q rewrite multiplicationNIsCommutative x 0 = refl
inject (negSucc x) = negativeQ record { numerator = succ x ; denominator = 1 ; denomPos = le zero refl ; hcf = record { c = 1 ; c|a = aDivA 1 ; c|b = oneDivN (succ x) ; hcf = λ i i|1 i|x → i|1 } ; coprime = le zero refl }
where
q : 1 ≡ x *N 0 +N 1
q rewrite multiplicationNIsCommutative x 0 = refl
cancel : (numerator denominator : ) → (0 <N denominator) → Abs
cancel num zero ()
cancel num (succ denom) _ with euclid num (succ denom)
cancel num (succ denom) _ | record { hcf = record { c = c ; c|a = divides record { quot = num/c ; rem = .0 ; pr = c*num/c ; remIsSmall = 0<c } refl ; c|b = divides record { quot = sd/c ; rem = .0 ; pr = c*d/c ; remIsSmall = 0<c' } refl ; hcf = hcf } } = record { numerator = num/c ; denominator = sd/c ; denomPos = 0<sd/c sd/c c*d/c ; hcf = h ; coprime = le zero refl }
where
lemm : {b c : } → (c ≡ 0) → (c *N b ≡ 0)
lemm {b} {c} c=0 rewrite c=0 = refl
cNonzero : (c ≡ 0) → False
cNonzero c=0 = naughtE (identityOfIndiscernablesLeft _ _ _ _≡_ c*d/c (applyEquality (_+N 0) (lemm c=0)))
0<sd/c : (sd/c : ) → (c *N sd/c +N 0 ≡ succ denom) → 0 <N sd/c
0<sd/c zero pr rewrite multiplicationNIsCommutative c zero = naughtE pr
0<sd/c (succ sd/c) pr = succIsPositive _
h : hcfData sd/c num/c
h = record
{ c = 1
; c|a = oneDivN _
; c|b = oneDivN _
; hcf = {!!}
}
productPos : {denomA denomB : } → (0 <N denomA) → (0 <N denomB) → 0 <N denomA *N denomB
productPos {zero} {b} ()
productPos {succ a} {zero} 0<a ()
productPos {succ a} {succ b} _ _ = succIsPositive _
addToPositive : (p : Abs) →
addToPositive p 0Q = positiveQ p
addToPositive (record { numerator = numA ; denominator = denomA ; denomPos = denomAPos ; hcf = hcfA ; coprime = coprimeA }) (positiveQ record { numerator = numB ; denominator = denomB ; denomPos = denomBPos ; hcf = hcfB ; coprime = coprimeB }) = positiveQ (cancel (numA *N denomB +N numB *N denomA) (denomA *N denomB) (productPos denomAPos denomBPos))
addToPositive (record { numerator = numA ; denominator = denomA ; denomPos = denomAPos ; hcf = hcfA ; coprime = coprimeA }) (negativeQ record { numerator = numB ; denominator = denomB ; denomPos = denomBPos ; hcf = hcfB ; coprime = coprimeB }) with orderIsTotal (numA *N denomB) (numB *N denomA)
... | inl (inl (le x pr)) = negativeQ (cancel x (denomA *N denomB) (productPos denomAPos denomBPos))
... | inl (inr (le x pr)) = positiveQ (cancel x (denomA *N denomB) (productPos denomAPos denomBPos))
... | inr x = 0Q
infix 15 _+Q_
_+Q_ :
0Q +Q b = b
positiveQ x +Q b = addToPositive x b
negativeQ a +Q 0Q = negativeQ a
negativeQ a +Q positiveQ x = addToPositive x (negativeQ a)
negativeQ record { numerator = numA ; denominator = denomA ; denomPos = denomAPos ; hcf = hcfA ; coprime = coprimeA } +Q negativeQ record { numerator = numB ; denominator = denomB ; denomPos = denomBPos ; hcf = hcfB ; coprime = coprimeB } = negativeQ (cancel (numA *N denomB +N numB *N denomA) (denomA *N denomB) (productPos denomAPos denomBPos))
negateQ :
negateQ 0Q = 0Q
negateQ (positiveQ x) = negativeQ x
negateQ (negativeQ x) = positiveQ x
negatePlusLeft : (a : ) → ((negateQ a) +Q a ≡ 0Q)
negatePlusLeft 0Q = refl
negatePlusLeft (positiveQ x) = {!!}
negatePlusLeft (negativeQ x) = {!!}
-}
{-
infix 25 _*Q_
_*Q_ :
record { numerator = numA ; denominator = zero ; denomPos = le x () } *Q b
record { numerator = numA ; denominator = succ denomA ; denomPos = denomPosA } *Q record { numerator = numB ; denominator = 0 ; denomPos = () }
record { sign = Positive ; numerator = numA ; denominator = succ denomA ; denomPos = denomPosA ; division = divisionA ; coprime = coprimeA } *Q record { sign = Positive ; numerator = numB ; denominator = succ denomB ; denomPos = denomPosB ; division = divisionB ; coprime = coprimeB } = record { sign = Positive ; numerator = {!!} ; denominator = {!!} ; denomPos = {!!} ; division = {!!} ; coprime = {!!} }
record { sign = Positive ; numerator = numA ; denominator = succ denomA ; denomPos = denomPosA ; division = divisionA ; coprime = coprimeA } *Q record { sign = Negative ; numerator = numB ; denominator = succ denomB ; denomPos = denomPosB ; division = divisionB ; coprime = coprimeB } = {!!}
record { sign = Negative ; numerator = numA ; denominator = succ denomA ; denomPos = denomPosA ; division = divisionA ; coprime = coprimeA } *Q record { sign = sgnB ; numerator = numB ; denominator = succ denomB ; denomPos = denomPosB ; division = divisionB ; coprime = coprimeB } = {!!}
-}

24
Reals.agda Normal file
View File

@@ -0,0 +1,24 @@
{-# OPTIONS --safe --warning=error #-}
open import Fields
open import Rings
open import Functions
open import Orders
open import LogicalFormulae
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
module Reals where
record Subset {m n : _} (A : Set m) (B : Set n) : Set (m n) where
field
inj : A B
injInj : Injection inj
record RealField {n : _} {A : Set n} {R : Ring A} {F : Field R} : Set _ where
field
orderedField : OrderedField F
open OrderedField orderedField
open TotalOrder ord
open Ring R
field
complete : {c : _} {C : Set c} (S : Subset C A) (upperBound : A) ({s : C} (Subset.inj S s) < upperBound) Sg A (λ lub ({s : C} (Subset.inj S s) < lub) && ({s : C} (Subset.inj S s) < upperBound (Subset.inj S s) < lub))

41
RedBlackTree.agda Normal file
View File

@@ -0,0 +1,41 @@
{-# OPTIONS --warning=error --safe #-}
open import LogicalFormulae
open import Functions
open import Maybe
open import Naturals
module RedBlackTree where
record BinaryTreeNode {a : _} (carrier : Set a) (order : TotalOrder carrier) (minValue : Maybe carrier) (maxValue : Maybe carrier) : Set a
valueExtractor : {a : _} {carrier : Set a} {order : TotalOrder carrier} {minValue : Maybe carrier} {maxValue : Maybe carrier} BinaryTreeNode {a} carrier order minValue maxValue carrier
record BinaryTreeNode {a} carrier order minValue maxValue where
inductive
field
value : carrier
min<=val : TotalOrder._<_ order min val
leftChild : Maybe (Sg (BinaryTreeNode {a} carrier order minValue (yes value)) (λ i TotalOrder._<_ order (valueExtractor {a} {carrier} {order} i) value))
rightChild : Maybe (Sg (BinaryTreeNode {a} carrier order (yes value) maxValue) (λ i TotalOrder._<_ order value (valueExtractor i)))
valueExtractor t = BinaryTreeNode.value t
lookup : {a : _} {carrier : Set a} {order : TotalOrder carrier} {minValue : Maybe carrier} {maxValue : Maybe carrier} (t : BinaryTreeNode carrier order minValue maxValue) (k : carrier) Maybe carrier
lookup {a} {carrier} {order} record { value = value ; leftChild = leftChild ; rightChild = rightChild } k with TotalOrder.totality order k value
lookup {a} {carrier} {order} record { value = value ; leftChild = no ; rightChild = rightChild } k | inl (inl k<val) = no
lookup {a} {carrier} {order} record { value = value ; leftChild = (yes (tree , _)) ; rightChild = rightChild } k | inl (inl k<val) = lookup tree k
lookup {a} {carrier} {order} record { value = value ; leftChild = leftChild ; rightChild = no } k | inl (inr val<k) = no
lookup {a} {carrier} {order} record { value = value ; leftChild = leftChild ; rightChild = (yes (tree , _)) } k | inl (inr val<k) = lookup tree k
lookup {a} {carrier} {order} record { value = value ; leftChild = leftChild ; rightChild = rightChild } k | inr k=val = yes value
addTree : {a : _} {carrier : Set a} {order : TotalOrder carrier} {minValue : Maybe carrier} {maxValue : Maybe carrier} (t : BinaryTreeNode carrier order minValue maxValue) (k : carrier) BinaryTreeNode carrier order (yes (defaultValue k minValue)) (yes (defaultValue k maxValue))
addTree {a} {carrier} {order} record { value = value ; leftChild = no ; rightChild = no } k with TotalOrder.totality order k value
addTree {a} {carrier} {order} record { value = value ; leftChild = no ; rightChild = no } k | inl (inr val<k) = record { value = value ; leftChild = no ; rightChild = yes (record { value = k ; leftChild = no ; rightChild = no} , val<k)}
addTree {a} {carrier} {order} record { value = value ; leftChild = no ; rightChild = no } k | inl (inl k<val) = record { value = value ; leftChild = yes (record { value = k ; leftChild = no ; rightChild = no} , k<val) ; rightChild = no }
addTree {a} {carrier} {order} record { value = value ; leftChild = no ; rightChild = no } k | inr k=val = record { value = k ; leftChild = no ; rightChild = no }
addTree {a} {carrier} {order} record { value = value ; leftChild = no ; rightChild = (yes x) } k with TotalOrder.totality order k value
addTree {a} {carrier} {order} record { value = value ; leftChild = no ; rightChild = (yes child) } k | inl (inl k<val) = {!!}
addTree {a} {carrier} {order} record { value = value ; leftChild = no ; rightChild = (yes child) } k | inl (inr val<k) = {!!}
addTree {a} {carrier} {order} {minValue} {maxValue} record { value = value ; leftChild = no ; rightChild = (yes (child , pr)) } k | inr k=val = record { value = k ; leftChild = no ; rightChild = yes (typeCast child (applyEquality (λ i BinaryTreeNode carrier order {!!} {!!}) {!!}) , {!!}) }
addTree {a} {carrier} {order} record { value = value ; leftChild = (yes x) ; rightChild = rightChild } k = {!!}

39
RingExamples.agda Normal file
View File

@@ -0,0 +1,39 @@
{-# OPTIONS --safe --warning=error #-}
open import LogicalFormulae
open import Groups
open import Functions
open import Naturals
open import Integers
open import IntegersModN
open import RingExamplesProofs
open import PrimeNumbers
module RingExamples where
nToZn : (n : ) (pr : 0 <N n) (x : ) n n pr
nToZn n pr x = nToZn' n pr x
mod : (n : ) (pr : 0 <N n) n n pr
mod n pr a = mod' n pr a
modNExampleSurjective : (n : ) (pr : 0 <N n) Surjection (mod n pr)
modNExampleSurjective n pr = modNExampleSurjective' n pr
{-
modNExampleGroupHom : (n : ) → (pr : 0 <N n) → GroupHom Group (nGroup n pr) (mod n pr)
modNExampleGroupHom n pr = modNExampleGroupHom' n pr
embedZnInZ : {n : } {pr : 0 <N n} → (a : n n pr) →
embedZnInZ record { x = x } = nonneg x
modNRoundTrip : (n : ) → (pr : 0 <N n) → (a : n n pr) → mod n pr (embedZnInZ a) ≡ a
modNRoundTrip zero ()
modNRoundTrip (succ n) pr record { x = x ; xLess = xLess } with divisionAlg (succ n) x
modNRoundTrip (succ n) _ record { x = x ; xLess = xLess } | record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = inl remIsSmall ; quotSmall = quotSmall } = equalityZn _ _ p
where
p : rem ≡ x
p = modIsUnique record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = inl remIsSmall ; quotSmall = quotSmall } record { quot = 0 ; rem = x ; pr = identityOfIndiscernablesLeft _ _ _ _≡_ refl (applyEquality (λ i → i +N x) (multiplicationNIsCommutative 0 n)) ; remIsSmall = inl xLess ; quotSmall = inl (succIsPositive n) }
modNRoundTrip (succ n) _ record { x = x ; xLess = xLess } | record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = inr () }
-}

135
RingExamplesProofs.agda Normal file
View File

@@ -0,0 +1,135 @@
{-# OPTIONS --safe --warning=error #-}
open import LogicalFormulae
open import Functions
open import Groups
open import Orders
open import Rings
open import Naturals
open import Integers
open import PrimeNumbers
open import IntegersModN
module RingExamplesProofs where
nToZn' : (n : ) (pr : 0 <N n) (x : ) n n pr
nToZn' 0 ()
nToZn' (succ n) pr x with divisionAlg (succ n) x
nToZn' (succ n) pr1 x | record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = inl thing ; quotSmall = quotSmall } = record { x = rem ; xLess = thing }
nToZn' (succ n) pr1 x | record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = inr () ; quotSmall = quotSmall }
mod' : (n : ) (pr : 0 <N n) n n pr
mod' zero () a
mod' (succ n) pr (nonneg x) = nToZn' (succ n) pr x
mod' (succ n) pr (negSucc x) = Group.inverse (nGroup (succ n) pr) (nToZn' (succ n) pr (succ x))
subtractionEquiv : (a : ) {b c : } (c<b : c <N b) a +N c b a subtractionNResult.result (-N (inl c<b))
subtractionEquiv 0 {b} {c} c<b pr rewrite pr = exFalso (PartialOrder.irreflexive (TotalOrder.order TotalOrder) c<b)
subtractionEquiv (succ a) {b} {c} c<b pr = equivalentSubtraction 0 b (succ a) c (succIsPositive a) c<b (equalityCommutative pr)
modNExampleSurjective' : (n : ) (pr : 0 <N n) Surjection (mod' n pr)
Surjection.property (modNExampleSurjective' zero ())
Surjection.property (modNExampleSurjective' (succ n) pr) record { x = x ; xLess = xLess } with divisionAlg (succ n) x
Surjection.property (modNExampleSurjective' (succ n) p) record { x = x ; xLess = xLess } | record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = inl remIsSmall ; quotSmall = q } = nonneg x , lhs'
where
rs' : rem x
rs' = modIsUnique (record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = inl remIsSmall ; quotSmall = q }) (record { quot = 0 ; rem = x ; pr = blah ; remIsSmall = inl xLess ; quotSmall = inl (succIsPositive n) })
where
blah : n *N 0 +N x x
blah rewrite multiplicationNIsCommutative n 0 = refl
lhs : nToZn' (succ n) p x record { x = rem ; xLess = remIsSmall }
lhs with divisionAlg (succ n) x
lhs | record { quot = quot' ; rem = rem' ; pr = pr' ; remIsSmall = inl t ; quotSmall = quotSmall } = equalityZn _ _ (equalityCommutative rs)
where
rs : rem rem'
rs = modIsUnique (record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = inl remIsSmall; quotSmall = q }) (record { quot = quot' ; rem = rem' ; pr = pr' ; remIsSmall = inl t ; quotSmall = quotSmall })
lhs | record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = inr () ; quotSmall = quotSmall }
lhs' : nToZn' (succ n) p x record { x = x ; xLess = xLess }
lhs' = transitivity lhs (equalityZn _ _ rs')
Surjection.property (modNExampleSurjective' (succ n) p) record { x = x ; xLess = xLess } | record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = inr () ; quotSmall = quotSmall }
{-
modNExampleGroupHom' : (n : ) → (pr : 0 <N n) → GroupHom Group (nGroup n pr) (mod' n pr)
modNExampleGroupHom' 0 ()
GroupHom.wellDefined (modNExampleGroupHom' (succ n) pr) {x} {.x} refl = refl
GroupHom.groupHom (modNExampleGroupHom' (succ n) _) {nonneg a} {nonneg b} with divisionAlg (succ n) a
GroupHom.groupHom (modNExampleGroupHom' (succ n) _) {nonneg a} {nonneg b} | record { quot = quotA ; rem = remA ; pr = prA ; remIsSmall = inl remA<sn ; quotSmall = quotSmallA } with divisionAlg (succ n) b
GroupHom.groupHom (modNExampleGroupHom' (succ n) _) {nonneg a} {nonneg b} | record { quot = quotA ; rem = remA ; pr = prA ; remIsSmall = inl remA<sn ; quotSmall = quotSmallA } | record { quot = quotB ; rem = remB ; pr = prB ; remIsSmall = inl remB<sn ; quotSmall = quotSmallB } with orderIsTotal (remA +N remB) (succ n)
GroupHom.groupHom (modNExampleGroupHom' (succ n) pr1) {nonneg a} {nonneg b} | record { quot = quotA ; rem = remA ; pr = prA ; remIsSmall = inl remA<sn ; quotSmall = _ } | record { quot = quotB ; rem = remB ; pr = prB ; remIsSmall = inl remB<sn ; quotSmall = _ } | inl (inl rarb<sn) rewrite addingNonnegIsHom a b = equalityZn _ _ lemma
where
lemma : n.x (nToZn' (succ n) pr1 (a +N b)) ≡ remA +N remB
lemma with divisionAlg (succ n) (a +N b)
lemma | record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = inl x ; quotSmall = inl _ } = equalityCommutative thing5
where
thing : ((succ n) *N quotA +N remA) +N ((succ n) *N quotB +N remB) ≡ a +N b
thing rewrite prA | prB = refl
thing2 : ((succ n) *N quotA +N remA) +N ((succ n) *N quotB +N remB) ≡ (succ n) *N quot +N rem
thing2 rewrite pr = thing
thing3 : (((succ n) *N quotA) +N ((succ n) *N quotB)) +N (remA +N remB) ≡ (succ n) *N quot +N rem
thing3 rewrite equalityCommutative (additionNIsAssociative (((succ n) *N quotA) +N ((succ n) *N quotB)) remA remB) | additionNIsAssociative ((succ n) *N quotA) ((succ n) *N quotB) remA | additionNIsCommutative ((succ n) *N quotB) remA | equalityCommutative (additionNIsAssociative ((succ n) *N quotA) remA ((succ n) *N quotB)) | additionNIsAssociative ((succ n) *N quotA +N remA) ((succ n) *N quotB) remB = thing2
thing4 : (succ n) *N (quotA +N quotB) +N (remA +N remB) ≡ (succ n) *N quot +N rem
thing4 rewrite productDistributes (succ n) quotA quotB = thing3
thing5 : remA +N remB ≡ rem
thing5 = modUniqueLemma {remA +N remB} {rem} {succ n} (quotA +N quotB) quot rarb<sn x thing4
lemma | record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = inl x ; quotSmall = inr () }
lemma | record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = inr () ; quotSmall = quotSmall }
GroupHom.groupHom (modNExampleGroupHom' (succ n) pr1) {nonneg a} {nonneg b} | record { quot = quotA ; rem = remA ; pr = prA ; remIsSmall = inl remA<sn } | record { quot = quotB ; rem = remB ; pr = prB ; remIsSmall = inl remB<sn } | inl (inr sn<rarb) rewrite addingNonnegIsHom a b = equalityZn _ _ lemma
where
lemma : n.x (nToZn' (succ n) pr1 (a +N b)) ≡ subtractionNResult.result (-N (inl sn<rarb))
lemma with divisionAlg (succ n) (a +N b)
lemma | record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = inl x ; quotSmall = q } = modIsUnique record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = inl x ; quotSmall = q } record { quot = succ quotA +N quotB ; rem = subtractionNResult.result (-N (inl sn<rarb)) ; pr = answer ; remIsSmall = inl remSmall ; quotSmall = inl (succIsPositive n) }
where
transform : (a : ) → {b c : } → (p : b <N c) → c <N a +N b → subtractionNResult.result (-N (inl p)) <N a
transform a {b} {c} pr (le y proof1) with addIntoSubtraction (succ y) {b} {c} (inl pr)
... | bl = le y (transitivity bl (equalityCommutative (subtractionEquiv a (orderIsTransitive pr (addingIncreases c y)) (equalityCommutative (identityOfIndiscernablesLeft _ _ _ _≡_ proof1 (additionNIsCommutative (succ y) c))))))
thing : ((succ n) *N quotA +N remA) +N ((succ n) *N quotB +N remB) ≡ a +N b
thing rewrite prA | prB = refl
thing2 : (((succ n) *N quotA) +N ((succ n) *N quotB)) +N (remA +N remB) ≡ a +N b
thing2 = identityOfIndiscernablesLeft _ _ _ _≡_ thing (transitivity (equalityCommutative (additionNIsAssociative ((quotA +N n *N quotA) +N remA) (succ n *N quotB) remB)) (transitivity (applyEquality (λ i → i +N remB) (additionNIsAssociative (quotA +N n *N quotA) remA (quotB +N n *N quotB))) (transitivity (applyEquality (λ i → ((quotA +N n *N quotA) +N i) +N remB) (additionNIsCommutative remA (quotB +N n *N quotB))) (transitivity (applyEquality (λ i → i +N remB) (equalityCommutative (additionNIsAssociative (quotA +N n *N quotA) (quotB +N n *N quotB) remA))) (additionNIsAssociative _ remA remB)))))
thing3 : (succ n) *N (quotA +N quotB) +N (remA +N remB) ≡ a +N b
thing3 = identityOfIndiscernablesLeft _ _ _ _≡_ thing2 (equalityCommutative (applyEquality (λ i → i +N (remA +N remB)) (productDistributes (succ n) (quotA) quotB)))
answer : (succ n *N succ (quotA +N quotB)) +N subtractionNResult.result (-N (inl sn<rarb)) ≡ a +N b
answer with addIntoSubtraction (succ n *N succ (quotA +N quotB)) (inl sn<rarb)
... | bl = transitivity bl (moveOneSubtraction' {a<=b = inl (orderIsTransitive sn<rarb (addingIncreases (remA +N remB) ((quotA +N quotB) +N n *N succ (quotA +N quotB))))} answer')
where
snTimes1 : succ n ≡ succ n *N 1
snTimes1 rewrite multiplicationNIsCommutative (succ n) 1 | additionNIsCommutative (succ n) 0 = refl
q' : succ n *N succ (quotA +N quotB) ≡ succ n +N (succ n *N (quotA +N quotB))
q' rewrite additionNIsCommutative (succ n) (succ n *N (quotA +N quotB)) | snTimes1 | equalityCommutative (productDistributes (succ n) (quotA +N quotB) 1) = applyEquality (λ i → (succ n) *N i) (succIsAddOne (quotA +N quotB))
answer'' : (succ n *N succ (quotA +N quotB)) +N (remA +N remB) ≡ (succ n) +N ((succ n *N (quotA +N quotB)) +N (remA +N remB))
answer'' rewrite equalityCommutative (additionNIsAssociative (succ n) (succ n *N (quotA +N quotB)) (remA +N remB)) = applyEquality (λ i → i +N (remA +N remB)) q'
answer' : (remA +N remB) +N (succ n *N succ (quotA +N quotB)) ≡ succ n +N (a +N b)
answer' rewrite equalityCommutative thing3 = transitivity (additionNIsCommutative (remA +N remB) (succ n *N succ (quotA +N quotB))) answer''
remSmall : subtractionNResult.result (-N (inl sn<rarb)) <N succ n
remSmall = transform (succ n) sn<rarb (addStrongInequalities remA<sn remB<sn)
lemma | record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = inr () ; quotSmall = quotSmall }
GroupHom.groupHom (modNExampleGroupHom' (succ n) pr1) {nonneg a} {nonneg b} | record { quot = quotA ; rem = remA ; pr = prA ; remIsSmall = inl remA<sn } | record { quot = quotB ; rem = remB ; pr = prB ; remIsSmall = inl remB<sn } | inr rarb=sn rewrite addingNonnegIsHom a b = equalityZn _ _ lemma
where
lemma : n.x (nToZn' (succ n) pr1 (a +N b)) ≡ 0
lemma with divisionAlg (succ n) (a +N b)
lemma | record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = inl x } = equalityCommutative (modUniqueLemma ((quotA +N quotB) +N 1) quot pr1 x thing7)
where
thing : ((succ n) *N quotA +N remA) +N ((succ n) *N quotB +N remB) ≡ a +N b
thing rewrite prA | prB = refl
thing2 : ((succ n) *N quotA +N remA) +N ((succ n) *N quotB +N remB) ≡ (succ n) *N quot +N rem
thing2 rewrite pr = thing
thing3 : (((succ n) *N quotA) +N ((succ n) *N quotB)) +N (remA +N remB) ≡ (succ n) *N quot +N rem
thing3 rewrite equalityCommutative (additionNIsAssociative (((succ n) *N quotA) +N ((succ n) *N quotB)) remA remB) | additionNIsAssociative ((succ n) *N quotA) ((succ n) *N quotB) remA | additionNIsCommutative ((succ n) *N quotB) remA | equalityCommutative (additionNIsAssociative ((succ n) *N quotA) remA ((succ n) *N quotB)) | additionNIsAssociative ((succ n) *N quotA +N remA) ((succ n) *N quotB) remB = thing2
thing4 : (succ n) *N (quotA +N quotB) +N (remA +N remB) ≡ (succ n) *N quot +N rem
thing4 rewrite productDistributes (succ n) quotA quotB = thing3
thing5 : (succ n) *N (quotA +N quotB) +N (succ n) ≡ (succ n) *N quot +N rem
thing5 rewrite equalityCommutative rarb=sn = thing4
thing6 : (succ n) *N ((quotA +N quotB) +N 1) ≡ (succ n) *N quot +N rem
thing6 rewrite productDistributes (succ n) (quotA +N quotB) 1 | multiplicationNIsCommutative n 1 | additionNIsCommutative n 0 = thing5
thing7 : (succ n) *N ((quotA +N quotB) +N 1) +N 0 ≡ (succ n) *N quot +N rem
thing7 = identityOfIndiscernablesLeft _ _ _ _≡_ thing6 (additionNIsCommutative 0 _)
lemma | record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = inr () ; quotSmall = quotSmall }
GroupHom.groupHom (modNExampleGroupHom' (succ n) _) {nonneg a} {nonneg b} | record { quot = quotA ; rem = remA ; pr = prA ; remIsSmall = inl remA<sn ; quotSmall = quotSmallA } | record { quot = quotB ; rem = remB ; pr = prB ; remIsSmall = inr () ; quotSmall = quotSmallB }
GroupHom.groupHom (modNExampleGroupHom' (succ n) _) {nonneg a} {nonneg b} | record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = inr () ; quotSmall = quotSmall }
GroupHom.groupHom (modNExampleGroupHom' (succ n) _) {nonneg a} {negSucc b} with divisionAlg (succ n) a
GroupHom.groupHom (modNExampleGroupHom' (succ n) _) {nonneg a} {negSucc b} | record { quot = quotA ; rem = remA ; pr = prA ; remIsSmall = remIsSmallA ; quotSmall = quotSmallA } = {!!}
GroupHom.groupHom (modNExampleGroupHom' (succ n) _) {negSucc x} {nonneg b} with divisionAlg (succ n) (succ x)
... | bl = {!!}
GroupHom.groupHom (modNExampleGroupHom' (succ n) _) {negSucc x} {negSucc b} with divisionAlg (succ n) (succ x)
... | bl = {!!}
-}

82
Rings.agda Normal file
View File

@@ -0,0 +1,82 @@
{-# OPTIONS --safe --warning=error #-}
open import LogicalFormulae
open import Groups
open import Naturals
open import Orders
open import Setoids
open import Functions
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
-- Following Part IB's course Groups, Rings, and Modules, we take rings to be commutative with one.
module Rings where
record Ring {n m} {A : Set n} (S : Setoid {n} {m} A) (_+_ : A A A) (_*_ : A A A) : Set (lsuc n m) where
field
additiveGroup : Group S _+_
open Group additiveGroup
open Setoid S
0R : A
0R = identity
field
multWellDefined : {r s t u : A} (r t) (s u) r * s t * u
1R : A
groupIsAbelian : {a b : A} a + b b + a
multAssoc : {a b c : A} (a * (b * c)) (a * b) * c
multCommutative : {a b : A} a * b b * a
multDistributes : {a b c : A} a * (b + c) (a * b) + (a * c)
multIdentIsIdent : {a : A} 1R * a a
record OrderedRing {n m o} {A : Set n} (S : Setoid {n} {m} A) (_+_ : A A A) (_*_ : A A A) : Set (lsuc n m lsuc o) where
field
ring : Ring S _+_ _*_
open Ring ring
open Group additiveGroup
open Setoid S
field
order : TotalOrder {_} {o} A
_<_ = TotalOrder._<_ order
field
orderRespectsAddition : {a b : A} (a < b) (c : A) (a + c) < (b + c)
orderRespectsMultiplication : {a b : A} (0R < a) (0R < b) (0R < (a * b))
--directSumRing : {m n : _} {A : Set m} {B : Set n} (r : Ring A) (s : Ring B) Ring (A && B)
--Ring.additiveGroup (directSumRing r s) = directSumGroup (Ring.additiveGroup r) (Ring.additiveGroup s)
--Ring._*_ (directSumRing r s) (a ,, b) (c ,, d) = (Ring._*_ r a c) ,, Ring._*_ s b d
--Ring.multWellDefined (directSumRing r s) (a ,, b) (c ,, d) = Ring.multWellDefined r a c ,, Ring.multWellDefined s b d
--Ring.1R (directSumRing r s) = Ring.1R r ,, Ring.1R s
--Ring.groupIsAbelian (directSumRing r s) = Ring.groupIsAbelian r ,, Ring.groupIsAbelian s
--Ring.multAssoc (directSumRing r s) = Ring.multAssoc r ,, Ring.multAssoc s
--Ring.multCommutative (directSumRing r s) = Ring.multCommutative r ,, Ring.multCommutative s
--Ring.multDistributes (directSumRing r s) = Ring.multDistributes r ,, Ring.multDistributes s
--Ring.multIdentIsIdent (directSumRing r s) = Ring.multIdentIsIdent r ,, Ring.multIdentIsIdent s
record RingHom {m n o p : _} {A : Set m} {B : Set n} {SA : Setoid {m} {o} A} {SB : Setoid {n} {p} B} {_+A_ : A A A} {_*A_ : A A A} (R : Ring SA _+A_ _*A_) {_+B_ : B B B} {_*B_ : B B B} (S : Ring SB _+B_ _*B_) (f : A B) : Set (m n o p) where
open Ring S
open Group additiveGroup
open Setoid SB
field
preserves1 : f (Ring.1R R) 1R
ringHom : {r s : A} f (r *A s) (f r) *B (f s)
groupHom : GroupHom (Ring.additiveGroup R) additiveGroup f
--record RingIso {m n : _} {A : Set m} {B : Set n} (R : Ring A) (S : Ring B) (f : A → B) : Set (m ⊔ n) where
-- field
-- ringHom : RingHom R S f
-- bijective : Bijection f
ringTimesZero : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A A A} {_*_ : A A A} (R : Ring S _+_ _*_) {x : A} Setoid.__ S (x * (Ring.0R R)) (Ring.0R R)
ringTimesZero {S = S} {_+_ = _+_} {_*_ = _*_} R {x} = symmetric (transitive blah'' (transitive (Group.multAssoc additiveGroup) (transitive (wellDefined invLeft reflexive) multIdentLeft)))
where
open Ring R
open Group additiveGroup
open Setoid S
open Transitive (Equivalence.transitiveEq (Setoid.eq S))
open Reflexive (Equivalence.reflexiveEq (Setoid.eq S))
open Symmetric (Equivalence.symmetricEq (Setoid.eq S))
blah : (x * 0R) (x * 0R) + (x * 0R)
blah = transitive (multWellDefined reflexive (symmetric multIdentRight)) multDistributes
blah' : (inverse (x * 0R)) + (x * 0R) (inverse (x * 0R)) + ((x * 0R) + (x * 0R))
blah' = wellDefined reflexive blah
blah'' : 0R (inverse (x * 0R)) + ((x * 0R) + (x * 0R))
blah'' = transitive (symmetric invLeft) blah'

7
Scratch.agda Normal file
View File

@@ -0,0 +1,7 @@
open import LogicalFormulae
open import Naturals
module Scratch where
lem : {P Q : Set} (P || (P False)) (P && Q False) Q P
lem {P} {Q} (inl x) pr q = exFalso (pr (record { fst = x ; snd = q }))
lem {P} {Q} (inr notP) pr q = {!!}

131
Setoids.agda Normal file
View File

@@ -0,0 +1,131 @@
{-# OPTIONS --safe --warning=error #-}
open import Functions
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import LogicalFormulae
module Setoids where
record Setoid {a} {b} (A : Set a) : Set (a lsuc b) where
infix 1 __
field
__ : A A Set b
eq : Equivalence __
open Equivalence eq
open Reflexive reflexiveEq
open Transitive transitiveEq
open Symmetric symmetricEq
~refl : {r : A} (r r)
~refl {r} = reflexive {r}
record Quotient {a} {b} {c} {A : Set a} {image : Set b} (S : Setoid {a} {c} A) : Set (a b c) where
open Setoid S
field
map : A image
mapWellDefined : {x y : A} x y map x map y
mapSurjective : Surjection map
mapInjective : {x y : A} map x map y x y
record SetoidToSet {a b c : _} {A : Set a} (S : Setoid {a} {c} A) (quotient : Set b) : Set (a b c) where
open Setoid S
field
project : A quotient
wellDefined : (x y : A) (x y) project x project y
surj : Surjection project
inj : (x y : A) project x project y x y
open Setoid
reflSetoid : {a : _} (A : Set a) Setoid A
__ (reflSetoid A) a b = a b
Reflexive.reflexive (Equivalence.reflexiveEq (eq (reflSetoid A))) = refl
Symmetric.symmetric (Equivalence.symmetricEq (eq (reflSetoid A))) {b} {.b} refl = refl
Transitive.transitive (Equivalence.transitiveEq (eq (reflSetoid A))) {b} {.b} {.b} refl refl = refl
directSumSetoid : {m n o p : _} {A : Set m} {B : Set n} (r : Setoid {m} {o} A) (s : Setoid {n} {p} B) Setoid (A && B)
__ (directSumSetoid r s) (a ,, b) (c ,, d) = (Setoid.__ r a c) && (Setoid.__ s b d)
Reflexive.reflexive (Equivalence.reflexiveEq (eq (directSumSetoid r s))) {(a ,, b)} = Reflexive.reflexive (Equivalence.reflexiveEq (Setoid.eq r)) ,, Reflexive.reflexive (Equivalence.reflexiveEq (Setoid.eq s))
Symmetric.symmetric (Equivalence.symmetricEq (eq (directSumSetoid r s))) {(a ,, b)} {(c ,, d)} (fst ,, snd) = Symmetric.symmetric (Equivalence.symmetricEq (Setoid.eq r)) fst ,, Symmetric.symmetric (Equivalence.symmetricEq (Setoid.eq s)) snd
Transitive.transitive (Equivalence.transitiveEq (eq (directSumSetoid r s))) {a ,, b} {c ,, d} {e ,, f} (fst1 ,, snd1) (fst2 ,, snd2) = Transitive.transitive (Equivalence.transitiveEq (Setoid.eq r)) fst1 fst2 ,, Transitive.transitive (Equivalence.transitiveEq (Setoid.eq s)) snd1 snd2
record SetoidInjection {a b c d : _} {A : Set a} {B : Set b} (S : Setoid {a} {c} A) (T : Setoid {b} {d} B) (f : A B) : Set (a b c d) where
open Setoid S renaming (__ to _A_)
open Setoid T renaming (__ to _B_)
field
wellDefined : {x y : A} x A y (f x) B (f y)
injective : {x y : A} (f x) B (f y) x A y
record SetoidSurjection {a b c d : _} {A : Set a} {B : Set b} (S : Setoid {a} {c} A) (T : Setoid {b} {d} B) (f : A B) : Set (a b c d) where
open Setoid S renaming (__ to _A_)
open Setoid T renaming (__ to _B_)
field
wellDefined : {x y : A} x A y (f x) B (f y)
surjective : {x : B} Sg A (λ a f a B x)
record SetoidBijection {a b c d : _} {A : Set a} {B : Set b} (S : Setoid {a} {c} A) (T : Setoid {b} {d} B) (f : A B) : Set (a b c d) where
field
inj : SetoidInjection S T f
surj : SetoidSurjection S T f
record SetoidsBiject {a b c d : _} {A : Set a} {B : Set b} (S : Setoid {a} {c} A) (T : Setoid {b} {d} B) : Set (a b c d) where
field
bij : A B
bijIsBijective : SetoidBijection S T bij
setoidInjComp : {a b c d e f : _} {A : Set a} {B : Set b} {C : Set c} {S : Setoid {a} {d} A} {T : Setoid {b} {e} B} {U : Setoid {c} {f} C} {g : A B} {h : B C} (gB : SetoidInjection S T g) (hB : SetoidInjection T U h) SetoidInjection S U (h g)
SetoidInjection.wellDefined (setoidInjComp gI hI) x~y = SetoidInjection.wellDefined hI (SetoidInjection.wellDefined gI x~y)
SetoidInjection.injective (setoidInjComp gI hI) hgx~hgy = SetoidInjection.injective gI (SetoidInjection.injective hI hgx~hgy)
setoidSurjComp : {a b c d e f : _} {A : Set a} {B : Set b} {C : Set c} {S : Setoid {a} {d} A} {T : Setoid {b} {e} B} {U : Setoid {c} {f} C} {g : A B} {h : B C} (gB : SetoidSurjection S T g) (hB : SetoidSurjection T U h) SetoidSurjection S U (h g)
SetoidSurjection.wellDefined (setoidSurjComp gI hI) x~y = SetoidSurjection.wellDefined hI (SetoidSurjection.wellDefined gI x~y)
SetoidSurjection.surjective (setoidSurjComp gI hI) {x} with SetoidSurjection.surjective hI {x}
SetoidSurjection.surjective (setoidSurjComp gI hI) {x} | a , prA with SetoidSurjection.surjective gI {a}
SetoidSurjection.surjective (setoidSurjComp {U = U} gI hI) {x} | a , prA | b , prB = b , transitive (SetoidSurjection.wellDefined hI prB) prA
where
open Setoid U
open Transitive (Equivalence.transitiveEq (Setoid.eq U))
setoidBijComp : {a b c d e f : _} {A : Set a} {B : Set b} {C : Set c} {S : Setoid {a} {d} A} {T : Setoid {b} {e} B} {U : Setoid {c} {f} C} {g : A B} {h : B C} (gB : SetoidBijection S T g) (hB : SetoidBijection T U h) SetoidBijection S U (h g)
SetoidBijection.inj (setoidBijComp gB hB) = setoidInjComp (SetoidBijection.inj gB) (SetoidBijection.inj hB)
SetoidBijection.surj (setoidBijComp gB hB) = setoidSurjComp (SetoidBijection.surj gB) (SetoidBijection.surj hB)
setoidIdIsBijective : {a b : _} {A : Set a} {S : Setoid {a} {b} A} SetoidBijection S S (λ i i)
SetoidInjection.wellDefined (SetoidBijection.inj (setoidIdIsBijective {A = A})) = id
SetoidInjection.injective (SetoidBijection.inj (setoidIdIsBijective {A = A})) = id
SetoidSurjection.wellDefined (SetoidBijection.surj (setoidIdIsBijective {A = A})) = id
SetoidSurjection.surjective (SetoidBijection.surj (setoidIdIsBijective {S = S})) {x} = x , Reflexive.reflexive (Equivalence.reflexiveEq (Setoid.eq S))
record SetoidInvertible {a b c d : _} {A : Set a} {B : Set b} (S : Setoid {a} {c} A) (T : Setoid {b} {d} B) (f : A B) : Set (a b c d) where
field
fWellDefined : {x y : A} Setoid.__ S x y Setoid.__ T (f x) (f y)
inverse : B A
inverseWellDefined : {x y : B} Setoid.__ T x y Setoid.__ S (inverse x) (inverse y)
isLeft : (b : B) Setoid.__ T (f (inverse b)) b
isRight : (a : A) Setoid.__ S (inverse (f a)) a
setoidBijectiveImpliesInvertible : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {f : A B} (bij : SetoidBijection S T f) SetoidInvertible S T f
SetoidInvertible.fWellDefined (setoidBijectiveImpliesInvertible bij) x~y = SetoidInjection.wellDefined (SetoidBijection.inj bij) x~y
SetoidInvertible.inverse (setoidBijectiveImpliesInvertible bij) x with SetoidSurjection.surjective (SetoidBijection.surj bij) {x}
SetoidInvertible.inverse (setoidBijectiveImpliesInvertible bij) x | a , b = a
SetoidInvertible.inverseWellDefined (setoidBijectiveImpliesInvertible bij) {x} {y} x~y with SetoidSurjection.surjective (SetoidBijection.surj bij) {x}
SetoidInvertible.inverseWellDefined (setoidBijectiveImpliesInvertible {T = T} bij) {x} {y} x~y | a , prA with SetoidSurjection.surjective (SetoidBijection.surj bij) {y}
... | b , prB = SetoidInjection.injective (SetoidBijection.inj bij) (transitive prA (transitive x~y (symmetric prB)))
where
open Setoid T
open Transitive (Equivalence.transitiveEq (Setoid.eq T))
open Symmetric (Equivalence.symmetricEq (Setoid.eq T))
SetoidInvertible.isLeft (setoidBijectiveImpliesInvertible bij) b with SetoidSurjection.surjective (SetoidBijection.surj bij) {b}
... | a , prA = prA
SetoidInvertible.isRight (setoidBijectiveImpliesInvertible {f = f} bij) b with SetoidSurjection.surjective (SetoidBijection.surj bij) {f b}
... | fb , prFb = SetoidInjection.injective (SetoidBijection.inj bij) prFb
setoidInvertibleImpliesBijective : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {f : A B} (inv : SetoidInvertible S T f) SetoidBijection S T f
SetoidInjection.wellDefined (SetoidBijection.inj (setoidInvertibleImpliesBijective inv)) x~y = SetoidInvertible.fWellDefined inv x~y
SetoidInjection.injective (SetoidBijection.inj (setoidInvertibleImpliesBijective {S = S} {f = f} inv)) {x} {y} pr = transitive (symmetric (SetoidInvertible.isRight inv x)) (transitive (SetoidInvertible.inverseWellDefined inv pr) (SetoidInvertible.isRight inv y))
where
open Setoid S
open Transitive (Equivalence.transitiveEq (Setoid.eq S))
open Symmetric (Equivalence.symmetricEq (Setoid.eq S))
SetoidSurjection.wellDefined (SetoidBijection.surj (setoidInvertibleImpliesBijective inv)) x~y = SetoidInvertible.fWellDefined inv x~y
SetoidSurjection.surjective (SetoidBijection.surj (setoidInvertibleImpliesBijective inv)) {x} = SetoidInvertible.inverse inv x , SetoidInvertible.isLeft inv x

61
SymmetryGroups.agda Normal file
View File

@@ -0,0 +1,61 @@
{-# OPTIONS --safe --warning=error #-}
open import LogicalFormulae
open import Setoids
open import Functions
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Naturals
open import FinSet
open import Groups
open import GroupActions
module SymmetryGroups where
data SymmetryGroupElements {a b : _} {A : Set a} (S : Setoid {a} {b} A) : Set (a b) where
sym : {f : A A} SetoidBijection S S f SymmetryGroupElements S
data ExtensionallyEqual {a b : _} {A : Set a} {B : Set b} (f g : A B) : Set (a b) where
eq : ({x : A} (f x) (g x)) ExtensionallyEqual f g
extensionallyEqualReflexive : {a b : _} {A : Set a} {B : Set b} (f : A B) ExtensionallyEqual f f
extensionallyEqualReflexive f = eq (λ {x} refl)
extensionallyEqualSymmetric : {a b : _} {A : Set a} {B : Set b} {f g : A B} ExtensionallyEqual f g ExtensionallyEqual g f
extensionallyEqualSymmetric {f} {g} (eq pr) = eq λ {x} equalityCommutative (pr {x})
extensionallyEqualTransitive : {a b : _} {A : Set a} {B : Set b} {f g h : A B} ExtensionallyEqual f g ExtensionallyEqual g h ExtensionallyEqual f h
extensionallyEqualTransitive (eq pr1) (eq pr2) = eq λ {x} transitivity pr1 pr2
symmetricSetoid : {a b : _} {A : Set a} (S : Setoid {a} {b} A) Setoid (SymmetryGroupElements S)
Setoid.__ (symmetricSetoid A) (sym {f} bijF) (sym {g} bijG) = ExtensionallyEqual f g
Reflexive.reflexive (Equivalence.reflexiveEq (Setoid.eq (symmetricSetoid A))) {sym {f} bijF} = extensionallyEqualReflexive f
Symmetric.symmetric (Equivalence.symmetricEq (Setoid.eq (symmetricSetoid A))) {sym {f} bijF} {sym {g} bijG} f~g = extensionallyEqualSymmetric f~g
Transitive.transitive (Equivalence.transitiveEq (Setoid.eq (symmetricSetoid A))) {sym {f} bijF} {sym {g} bijG} {sym {h} bijH} f~g g~h = extensionallyEqualTransitive f~g g~h
symmetricGroupOp : {a b : _} {A : Set a} {S : Setoid {a} {b} A} (f g : SymmetryGroupElements S) SymmetryGroupElements S
symmetricGroupOp (sym {f} bijF) (sym {g} bijG) = sym (setoidBijComp bijF bijG)
symmetricGroupInv : {a b : _} {A : Set a} (S : Setoid {a} {b} A) SymmetryGroupElements S SymmetryGroupElements S
symmetricGroupInv S (sym {f} bijF) with setoidBijectiveImpliesInvertible bijF
... | record { inverse = inverse ; inverseWellDefined = iwd ; isLeft = isLeft ; isRight = isRight } = sym (setoidInvertibleImpliesBijective (record { fWellDefined = iwd ; inverse = f ; inverseWellDefined = SetoidInjection.wellDefined (SetoidBijection.inj bijF) ; isLeft = λ b isRight b ; isRight = λ b isLeft b }))
symmetricGroupInvIsLeft : {a b : _} {A : Set a} (S : Setoid {a} {b} A) {x : SymmetryGroupElements S} Setoid.__ (symmetricSetoid S) (symmetricGroupOp (symmetricGroupInv S x) x) (sym setoidIdIsBijective)
symmetricGroupInvIsLeft {A = A} S {sym fBij} = ExtensionallyEqual.eq λ {x} {!ans (sym fBij)!}
where
ans : (f : A A) (bij : SetoidBijection S S f) f (symmetricGroupInv S (sym fBij) x) sym fBij
ans elt = {!!}
symmetricGroupInvIsRight : {a b : _} {A : Set a} (S : Setoid {a} {b} A) {x : SymmetryGroupElements S} Setoid.__ (symmetricSetoid S) (symmetricGroupOp x (symmetricGroupInv S x)) (sym setoidIdIsBijective)
symmetricGroupInvIsRight S {sym x} = {!!}
symmetricGroup : {a b : _} {A : Set a} (S : Setoid {a} {b} A) Group (symmetricSetoid S) (symmetricGroupOp {A = A})
Group.wellDefined (symmetricGroup A) {sym {m} bijM} {sym {n} bijN} {sym {x} bijX} {sym {y} bijY} (ExtensionallyEqual.eq m~x) (ExtensionallyEqual.eq n~y) = ExtensionallyEqual.eq λ {z} transitivity (applyEquality n m~x) n~y
Group.identity (symmetricGroup A) = sym setoidIdIsBijective
Group.inverse (symmetricGroup S) = symmetricGroupInv S
Group.multAssoc (symmetricGroup A) {sym {f} bijF} {sym {g} bijG} {sym {h} bijH} = ExtensionallyEqual.eq λ {x} refl
Group.multIdentRight (symmetricGroup A) {sym {f} bijF} = ExtensionallyEqual.eq λ {x} refl
Group.multIdentLeft (symmetricGroup A) {sym {f} bijF} = ExtensionallyEqual.eq λ {x} refl
Group.invLeft (symmetricGroup S) {x} = symmetricGroupInvIsLeft S {x}
Group.invRight (symmetricGroup S) {x} = symmetricGroupInvIsRight S {x}
actionInducesHom : {a b c d : _} {A : Set a} {S : Setoid {a} {c} A} {_+_ : A A A} {G : Group S _+_} {B : Set b} {X : Setoid {b} {d} B} (GroupAction G X) SymmetryGroupElements X
actionInducesHom = {!!}

779
TempIntegers.agda Normal file
View File

@@ -0,0 +1,779 @@
{-# OPTIONS --safe --warning=error #-}
open import LogicalFormulae
open import Naturals
open import Groups
open import Rings
open import Integers
module TempIntegers where
additiveInverseExists : (a : ) (negSucc a +Z nonneg (succ a)) nonneg 0
additiveInverseExists zero = refl
additiveInverseExists (succ a) = additiveInverseExists a
multiplicationZIsCommutativeNonnegNegsucc : (a b : ) (nonneg a *Z negSucc b) negSucc b *Z nonneg a
multiplicationZIsCommutativeNonnegNegsucc zero zero = refl
multiplicationZIsCommutativeNonnegNegsucc zero (succ b) = refl
multiplicationZIsCommutativeNonnegNegsucc (succ a) zero = refl
multiplicationZIsCommutativeNonnegNegsucc (succ x) (succ y) = t
where
r : (nonneg (succ x) *Z negSucc y) (negSucc y) *Z (nonneg (succ x))
r = multiplicationZIsCommutativeNonnegNegsucc (succ x) y
s : negSucc x +Z (nonneg (succ x) *Z negSucc y) negSucc x +Z negSucc y *Z (nonneg (succ x))
s = applyEquality (_+Z_ (negSucc x)) r
t : negSucc x +Z (nonneg (succ x) *Z negSucc y) negSucc y *Z (nonneg (succ x)) +Z negSucc x
t = identityOfIndiscernablesRight (negSucc x +Z (nonneg (succ x) *Z negSucc y)) (negSucc x +Z negSucc y *Z (nonneg (succ x))) (negSucc y *Z nonneg (succ x) +Z negSucc x) _≡_ s (additionZIsCommutative (negSucc x) (negSucc y *Z nonneg (succ x)))
multiplicationZIsCommutative : (a b : ) (a *Z b) (b *Z a)
multiplicationZIsCommutative (nonneg x) (nonneg y) = applyEquality nonneg (multiplicationNIsCommutative x y)
multiplicationZIsCommutative (nonneg x) (negSucc y) = multiplicationZIsCommutativeNonnegNegsucc x y
multiplicationZIsCommutative (negSucc x) (nonneg y) = equalityCommutative (multiplicationZIsCommutativeNonnegNegsucc y x)
multiplicationZIsCommutative (negSucc zero) (negSucc y) = applyEquality nonneg (applyEquality succ n)
where
k : succ zero *N succ y succ y *N succ zero
k = multiplicationNIsCommutative (succ zero) (succ y)
j : succ y +N zero *N succ y succ zero +N y *N succ zero
j = transitivity refl (transitivity k refl)
alterL : y +N succ (zero *N succ y) succ zero +N y *N succ zero
alterL = identityOfIndiscernablesLeft (succ y +N zero *N succ y) (succ zero +N y *N succ zero) (y +N succ (zero *N succ y)) _≡_ j (equalityCommutative (succCanMove y (zero *N succ y)))
j' : y +N succ (zero *N succ y) zero +N succ (y *N succ zero)
j' = identityOfIndiscernablesRight (y +N succ (zero *N succ y)) (succ zero +N y *N succ zero) (zero +N succ (y *N succ zero)) _≡_ alterL (equalityCommutative (succCanMove zero (y *N succ zero)))
m : succ (y +N zero *N succ y) zero +N succ (y *N succ zero)
m = identityOfIndiscernablesLeft (y +N succ (zero *N succ y)) (zero +N succ (y *N succ zero)) (succ (y +N zero *N succ y)) _≡_ j' (succExtracts y (zero *N succ y))
m' : succ (y +N zero *N succ y) succ (zero +N y *N succ zero)
m' = identityOfIndiscernablesRight (succ (y +N zero *N succ y)) (zero +N succ (y *N succ zero)) (succ (zero +N y *N succ zero)) _≡_ m (succExtracts zero (y *N succ zero))
n : y +N zero *N succ y zero +N y *N succ zero
n = succInjective m'
multiplicationZIsCommutative (negSucc (succ x)) (negSucc y) = applyEquality nonneg (applyEquality succ n)
where
k : succ (succ x) *N succ y succ y *N succ (succ x)
k = multiplicationNIsCommutative (succ (succ x)) (succ y)
j : succ y +N (succ x) *N succ y succ (succ x) +N y *N succ (succ x)
j = transitivity refl (transitivity k refl)
alterL : y +N succ ((succ x) *N succ y) succ (succ x) +N y *N succ (succ x)
alterL = identityOfIndiscernablesLeft (succ y +N (succ x) *N succ y) (succ (succ x) +N y *N succ (succ x)) (y +N succ ((succ x) *N succ y)) _≡_ j (equalityCommutative (succCanMove y ((succ x) *N succ y)))
j' : y +N succ ((succ x) *N succ y) (succ x) +N succ (y *N succ (succ x))
j' = identityOfIndiscernablesRight (y +N succ ((succ x) *N succ y)) (succ (succ x) +N y *N succ (succ x)) ((succ x) +N succ (y *N succ (succ x))) _≡_ alterL (equalityCommutative (succCanMove (succ x) (y *N succ (succ x))))
m : succ (y +N (succ x) *N succ y) (succ x) +N succ (y *N succ (succ x))
m = identityOfIndiscernablesLeft (y +N succ ((succ x) *N succ y)) ((succ x) +N succ (y *N succ (succ x))) (succ (y +N (succ x) *N succ y)) _≡_ j' (succExtracts y ((succ x) *N succ y))
m' : succ (y +N (succ x) *N succ y) succ ((succ x) +N y *N succ (succ x))
m' = identityOfIndiscernablesRight (succ (y +N (succ x) *N succ y)) ((succ x) +N succ (y *N succ (succ x))) (succ ((succ x) +N y *N succ (succ x))) _≡_ m (succExtracts (succ x) (y *N succ (succ x)))
n : y +N (succ x) *N succ y (succ x) +N y *N succ (succ x)
n = succInjective m'
negSuccIsNotNonneg : (a b : ) (negSucc a nonneg b) False
negSuccIsNotNonneg zero b ()
negSuccIsNotNonneg (succ a) b ()
negSuccInjective : {a b : } (negSucc a negSucc b) a b
negSuccInjective refl = refl
zeroMultInZLeft : (a : ) (nonneg zero) *Z a (nonneg zero)
zeroMultInZLeft a = identityOfIndiscernablesLeft (a *Z nonneg zero) (nonneg zero) (nonneg zero *Z a) _≡_ (zeroMultInZRight a) (multiplicationZIsCommutative a (nonneg zero))
additionZeroDoesNothing : (a : ) (a +Z nonneg zero a)
additionZeroDoesNothing (nonneg zero) = refl
additionZeroDoesNothing (nonneg (succ x)) = applyEquality nonneg (applyEquality succ (addZeroRight x))
additionZeroDoesNothing (negSucc x) = refl
canSubtractOne : (a : ) (b : ) negSucc zero +Z a negSucc (succ b) a negSucc b
canSubtractOne (nonneg zero) b pr = i
where
simplified : negSucc zero negSucc (succ b)
simplified = identityOfIndiscernablesLeft (negSucc zero +Z nonneg zero) (negSucc (succ b)) (negSucc zero) _≡_ pr refl
removeNegsucc : zero succ b
removeNegsucc = negSuccInjective simplified
f : False
f = succIsNonzero (equalityCommutative removeNegsucc)
i : (nonneg zero) negSucc b
i = exFalso f
canSubtractOne (nonneg (succ x)) b pr = exFalso (negSuccIsNotNonneg (succ b) x (equalityCommutative pr))
canSubtractOne (negSucc x) .x refl = refl
canAddOne : (a : ) (b : ) a negSucc b negSucc zero +Z a negSucc (succ b)
canAddOne (nonneg x) b ()
canAddOne (negSucc x) .x refl = refl
canAddOneReversed : (a : ) (b : ) a negSucc b a +Z negSucc zero negSucc (succ b)
canAddOneReversed a b pr = identityOfIndiscernablesLeft (negSucc zero +Z a) (negSucc (succ b)) (a +Z negSucc zero) _≡_ (canAddOne a b pr) (additionZIsCommutative (negSucc zero) a)
addToNegativeSelf : (a : ) negSucc a +Z nonneg a negSucc zero
addToNegativeSelf zero = refl
addToNegativeSelf (succ a) = addToNegativeSelf a
multiplicativeIdentityOneZNegsucc : (a : ) (negSucc a *Z nonneg (succ zero) negSucc a)
multiplicativeIdentityOneZNegsucc zero = refl
multiplicativeIdentityOneZNegsucc (succ x) = i
where
inter : negSucc x *Z nonneg (succ zero) negSucc x
i : negSucc x *Z nonneg (succ zero) +Z negSucc zero negSucc (succ x)
inter = multiplicativeIdentityOneZNegsucc x
h = applyEquality (λ x x +Z negSucc zero) inter
i = identityOfIndiscernablesRight (negSucc x *Z nonneg (succ zero) +Z negSucc zero) (negSucc x +Z negSucc zero) (negSucc (succ x)) _≡_ h (canAddOneReversed (negSucc x) x refl)
multiplicativeIdentityOneZ : (a : ) (a *Z nonneg (succ zero) a)
multiplicativeIdentityOneZ (nonneg x) = applyEquality nonneg (productWithOneRight x)
multiplicativeIdentityOneZ (negSucc x) = multiplicativeIdentityOneZNegsucc (x)
multiplicativeIdentityOneZLeft : (a : ) (nonneg (succ zero)) *Z a a
multiplicativeIdentityOneZLeft a = identityOfIndiscernablesLeft (a *Z nonneg (succ zero)) a (nonneg (succ zero) *Z a) _≡_ (multiplicativeIdentityOneZ a) (multiplicationZIsCommutative a (nonneg (succ zero)))
-- Define the equivalence of negSucc + nonneg = nonneg with the corresponding statements of naturals
negSuccPlusNonnegNonneg : (a : ) (b : ) (c : ) (negSucc a +Z (nonneg b) nonneg c) b c +N succ a
negSuccPlusNonnegNonneg zero zero c ()
negSuccPlusNonnegNonneg zero (succ b) .b refl rewrite succExtracts b zero | addZeroRight b = refl
negSuccPlusNonnegNonneg (succ a) zero c ()
negSuccPlusNonnegNonneg (succ a) (succ b) c pr with negSuccPlusNonnegNonneg a b c pr
... | prev = identityOfIndiscernablesRight (succ b) (succ (c +N succ a)) (c +N succ (succ a)) _≡_ (applyEquality succ prev) (equalityCommutative (succExtracts c (succ a)))
negSuccPlusNonnegNonnegConverse : (a : ) (b : ) (c : ) (b c +N succ a) (negSucc a +Z (nonneg b) nonneg c)
negSuccPlusNonnegNonnegConverse zero zero c pr rewrite succExtracts c zero = naughtE pr
negSuccPlusNonnegNonnegConverse zero (succ b) c pr rewrite additionNIsCommutative c 1 with succInjective pr
... | pr' = applyEquality nonneg pr'
negSuccPlusNonnegNonnegConverse (succ a) zero c pr rewrite succExtracts c (succ a) = naughtE pr
negSuccPlusNonnegNonnegConverse (succ a) (succ b) c pr rewrite succExtracts c (succ a) with succInjective pr
... | pr' = negSuccPlusNonnegNonnegConverse a b c pr'
-- Define the equivalence of negSucc + nonneg = negSucc with the corresponding statements of naturals
negSuccPlusNonnegNegsucc : (a b c : ) (negSucc a +Z (nonneg b) negSucc c) c +N b a
negSuccPlusNonnegNegsucc zero zero .0 refl = refl
negSuccPlusNonnegNegsucc zero (succ b) c ()
negSuccPlusNonnegNegsucc (succ a) zero .(succ a) refl rewrite addZeroRight a = refl
negSuccPlusNonnegNegsucc (succ a) (succ b) c pr with negSuccPlusNonnegNegsucc a b c pr
... | pr' = identityOfIndiscernablesLeft (succ (c +N b)) (succ a) (c +N succ b) _≡_ (applyEquality succ pr') (equalityCommutative (succExtracts c b))
negSuccPlusNonnegNegsuccConverse : (a b c : ) (c +N b a) (negSucc a +Z (nonneg b) negSucc c)
negSuccPlusNonnegNegsuccConverse zero zero c pr rewrite addZeroRight c = applyEquality negSucc (equalityCommutative pr)
negSuccPlusNonnegNegsuccConverse zero (succ b) c pr rewrite succExtracts c b = naughtE (equalityCommutative pr)
negSuccPlusNonnegNegsuccConverse (succ a) zero c pr rewrite addZeroRight c = applyEquality negSucc (equalityCommutative pr)
negSuccPlusNonnegNegsuccConverse (succ a) (succ b) c pr rewrite succExtracts c b = negSuccPlusNonnegNegsuccConverse a b c (succInjective pr)
-- Define the impossibility of negSucc + negSucc = nonneg
negSuccPlusNegSuccIsNegative : (a b c : ) ((negSucc a) +Z (negSucc b) nonneg c) False
negSuccPlusNegSuccIsNegative zero b c ()
negSuccPlusNegSuccIsNegative (succ a) b c ()
-- Define the equivalence of negSucc + negSucc = negSucc
negSuccPlusNegSuccNegSucc : (a b c : ) (negSucc a) +Z (negSucc b) (negSucc c) succ (a +N b) c
negSuccPlusNegSuccNegSucc zero b .(succ b) refl = refl
negSuccPlusNegSuccNegSucc (succ a) b .(succ (a +N succ b)) refl = applyEquality succ (equalityCommutative (succExtracts a b))
negSuccPlusNegSuccNegSuccConverse : (a b c : ) succ (a +N b) c (negSucc a) +Z (negSucc b) negSucc c
negSuccPlusNegSuccNegSuccConverse zero b c pr = applyEquality negSucc pr
negSuccPlusNegSuccNegSuccConverse (succ a) b zero ()
negSuccPlusNegSuccNegSuccConverse (succ a) b (succ c) pr rewrite succExtracts a b = applyEquality negSucc pr
-- Define the equivalence of nonneg + nonneg = nonneg
nonnegPlusNonnegNonneg : (a b c : ) nonneg a +Z nonneg b nonneg c a +N b c
nonnegPlusNonnegNonneg a b c pr rewrite addingNonnegIsHom a b = nonnegInj pr
where
nonnegInj : {x y : } (pr : nonneg x nonneg y) x y
nonnegInj {x} {.x} refl = refl
nonnegPlusNonnegNonnegConverse : (a b c : ) a +N b c nonneg a +Z nonneg b nonneg c
nonnegPlusNonnegNonnegConverse a b c pr rewrite addingNonnegIsHom a b = applyEquality nonneg pr
nonnegIsNotNegsucc : {x y : } nonneg x negSucc y False
nonnegIsNotNegsucc {x} {y} ()
-- Define the impossibility of nonneg + nonneg = negSucc
nonnegPlusNonnegNegsucc : (a b c : ) nonneg a +Z nonneg b negSucc c False
nonnegPlusNonnegNegsucc a b c pr rewrite addingNonnegIsHom a b = nonnegIsNotNegsucc pr
-- Move negSucc to other side of equation
moveNegsucc : (a : ) (b c : ) (negSucc a) +Z b c b c +Z (nonneg (succ a))
moveNegsucc a (nonneg x) (nonneg y) pr with (negSuccPlusNonnegNonneg a x y pr)
... | pr' = identityOfIndiscernablesRight (nonneg x) (nonneg (y +N succ a)) (nonneg y +Z nonneg (succ a)) _≡_ (applyEquality nonneg pr') (equalityCommutative (addingNonnegIsHom y (succ a)))
moveNegsucc a (nonneg x) (negSucc y) pr with negSuccPlusNonnegNegsucc a x y pr
... | pr' = equalityCommutative (negSuccPlusNonnegNonnegConverse y (succ a) x (g {a} {x} {y} (applyEquality succ pr')))
where
g : {a x y : } succ (y +N x) succ a succ a x +N succ y
g {a} {x} {y} p rewrite additionNIsCommutative y x | succExtracts x y = equalityCommutative p
moveNegsucc a (negSucc x) (nonneg y) pr = exFalso (negSuccPlusNegSuccIsNegative a x y pr)
moveNegsucc a (negSucc x) (negSucc y) pr with negSuccPlusNegSuccNegSucc a x y pr
... | pr' = equalityCommutative (negSuccPlusNonnegNegsuccConverse y (succ a) x g)
where
g : x +N succ a y
g rewrite succExtracts x a | additionNIsCommutative a x = pr'
moveNegsuccConverse : (a : ) (b c : ) (b c +Z (nonneg (succ a))) (negSucc a) +Z b c
moveNegsuccConverse zero (nonneg x) (nonneg y) pr with nonnegPlusNonnegNonneg y 1 x (equalityCommutative pr)
... | pr' rewrite (equalityCommutative (succIsAddOne y)) = g
where
g : negSucc 0 +Z nonneg x nonneg y
g rewrite equalityCommutative pr' = refl
moveNegsuccConverse zero (nonneg x) (negSucc y) pr with moveNegsucc y (nonneg 1) (nonneg x) (equalityCommutative pr)
... | pr' rewrite addingNonnegIsHom x (succ y) = g x y (nonnegInjective pr')
where
g : (x y : ) (1 (x +N succ y)) negSucc 0 +Z nonneg x negSucc y
g zero zero pr = refl
g zero (succ y) ()
g (succ x) y pr = naughtE (identityOfIndiscernablesRight zero (x +N succ y) (succ (x +N y)) _≡_ (succInjective pr) (succExtracts x y))
moveNegsuccConverse zero (negSucc x) (nonneg y) pr = exFalso (nonnegPlusNonnegNegsucc y 1 x (equalityCommutative pr))
moveNegsuccConverse zero (negSucc x) (negSucc y) pr with negSuccPlusNonnegNegsucc y 1 x (equalityCommutative pr)
... | pr' = applyEquality negSucc g
where
g : succ x y
g rewrite additionNIsCommutative x 1 = pr'
moveNegsuccConverse (succ a) (nonneg x) (nonneg y) pr with nonnegPlusNonnegNonneg y (succ (succ a)) x (equalityCommutative pr)
... | pr' = negSuccPlusNonnegNonnegConverse (succ a) x y (equalityCommutative pr')
moveNegsuccConverse (succ a) (nonneg x) (negSucc y) pr with negSuccPlusNonnegNonneg y (succ (succ a)) x (equalityCommutative pr)
... | pr' = negSuccPlusNonnegNegsuccConverse (succ a) x y (g a x y pr')
where
g : (a x y : ) succ (succ a) x +N succ y y +N x succ a
g a x y pr rewrite succExtracts x y | additionNIsCommutative x y = equalityCommutative (succInjective pr)
moveNegsuccConverse (succ a) (negSucc x) (nonneg z) pr = exFalso (nonnegPlusNonnegNegsucc z (succ (succ a)) x (equalityCommutative pr))
moveNegsuccConverse (succ a) (negSucc x) (negSucc z) pr with negSuccPlusNonnegNegsucc z (succ (succ a)) x (equalityCommutative pr)
... | pr' = applyEquality negSucc (g a x z pr')
where
g : (a x z : ) x +N succ (succ a) z succ (a +N succ x) z
g a x z pr rewrite succExtracts x (succ a) = identityOfIndiscernablesLeft (succ (x +N succ a)) z (succ (a +N succ x)) _≡_ pr (applyEquality succ (s x a))
where
s : (x a : ) x +N succ a a +N succ x
s x a rewrite succCanMove a x = additionNIsCommutative x (succ a)
-- By this point, any statement about addition can be moved from N into Z and from Z into N by applying moveNegSucc and its converse to eliminate any adding of negSucc.
sumOfNegsucc : (a b : ) (negSucc a +Z negSucc b) negSucc (succ (a +N b))
sumOfNegsucc a b = negSuccPlusNegSuccNegSuccConverse a b (succ (a +N b)) refl
additionZInjLemma : {a b c : } (nonneg a (nonneg a +Z nonneg b) +Z nonneg (succ c)) False
additionZInjLemma {a} {b} {c} pr = cannotAddKeepingEquality a (c +N b) pr2''
where
pr'' : nonneg a nonneg (a +N b) +Z nonneg (succ c)
pr'' = identityOfIndiscernablesRight (nonneg a) ((nonneg a +Z nonneg b) +Z nonneg (succ c)) (nonneg (a +N b) +Z nonneg (succ c)) _≡_ pr (applyEquality (λ i i +Z nonneg (succ c)) (addingNonnegIsHom a b))
pr''' : nonneg a nonneg ((a +N b) +N succ c)
pr''' = identityOfIndiscernablesRight (nonneg a) (nonneg (a +N b) +Z nonneg (succ c)) (nonneg ((a +N b) +N succ c)) _≡_ pr'' (addingNonnegIsHom (a +N b) (succ c))
pr2 : a (a +N b) +N succ c
pr2 = nonnegInjective pr'''
pr2' : a a +N (b +N succ c)
pr2' = identityOfIndiscernablesRight a ((a +N b) +N succ c) (a +N (b +N succ c)) _≡_ pr2 (additionNIsAssociative a b (succ c))
pr2'' : a a +N succ (c +N b)
pr2'' rewrite additionNIsCommutative (succ c) b = pr2'
additionZInjectiveFirstLemma : (a : ) (b c : ) (nonneg a +Z b nonneg a +Z c) (b c)
additionZInjectiveFirstLemma a (nonneg b) (nonneg c) pr rewrite (addingNonnegIsHom a b) | (addingNonnegIsHom a c) = applyEquality nonneg (canSubtractFromEqualityLeft {a} pr')
where
pr' : a +N b a +N c
pr' = nonnegInjective pr
additionZInjectiveFirstLemma a (nonneg b) (negSucc c) pr = exFalso (additionZInjLemma pr')
where
pr' : nonneg a (nonneg a +Z nonneg b) +Z nonneg (succ c)
pr' rewrite additionZIsCommutative (nonneg a) (negSucc c) = moveNegsucc c (nonneg a) (nonneg a +Z nonneg b) (equalityCommutative pr)
additionZInjectiveFirstLemma a (negSucc b) (nonneg x) pr = exFalso (additionZInjLemma pr')
where
pr' : nonneg a (nonneg a +Z nonneg x) +Z nonneg (succ b)
pr' rewrite additionZIsCommutative (nonneg a) (negSucc b) = moveNegsucc b (nonneg a) (nonneg a +Z nonneg x) pr
additionZInjectiveFirstLemma zero (negSucc zero) (negSucc x) pr = pr
additionZInjectiveFirstLemma zero (negSucc (succ b)) (negSucc x) pr = pr
additionZInjectiveFirstLemma (succ a) (negSucc zero) (negSucc x) pr = applyEquality negSucc (equalityCommutative qr'')
where
pr1 : negSucc x +Z nonneg (succ a) nonneg a
pr1 rewrite additionZIsCommutative (nonneg (succ a)) (negSucc x) = equalityCommutative pr
pr' : nonneg a +Z nonneg (succ x) nonneg (succ a)
pr' = equalityCommutative (moveNegsucc x (nonneg (succ a)) (nonneg a) pr1)
pr'' : nonneg (a +N succ x) nonneg (succ a)
pr'' rewrite equalityCommutative (addingNonnegIsHom a (succ x)) = pr'
pr''' : a +N succ x succ a
pr''' = nonnegInjective pr''
pr'''' : succ x +N a succ a
pr'''' rewrite additionNIsCommutative (succ x) a = pr'''
qr : succ (a +N x) succ a
qr rewrite additionNIsCommutative a x = pr''''
qr' : a +N x a +N 0
qr' rewrite addZeroRight a = succInjective qr
qr'' : x 0
qr'' = canSubtractFromEqualityLeft {a} qr'
additionZInjectiveFirstLemma (succ a) (negSucc (succ b)) (negSucc zero) pr = naughtE qr
where
pr' : nonneg a nonneg a +Z nonneg (succ b)
pr' rewrite additionZIsCommutative (nonneg a) (negSucc b) = moveNegsucc b (nonneg a) (nonneg a) pr
pr'' : nonneg a nonneg (a +N succ b)
pr'' rewrite equalityCommutative (addingNonnegIsHom a (succ b)) = pr'
pr''' : a +N 0 a +N succ b
pr''' rewrite addZeroRight a = nonnegInjective pr''
qr : 0 succ b
qr = canSubtractFromEqualityLeft {a} pr'''
additionZInjectiveFirstLemma (succ a) (negSucc (succ b)) (negSucc (succ x)) pr = go b x pr''
where
pr' : negSucc b negSucc x
pr' = additionZInjectiveFirstLemma a (negSucc b) (negSucc x) pr
pr'' : b x
pr'' = negSuccInjective pr'
go : (b x : ) (b x) negSucc (succ b) negSucc (succ x)
go b .b refl = refl
additionZInjectiveSecondLemmaLemma : (a : ) (b : ) (c : ) (negSucc a +Z nonneg b negSucc a +Z negSucc c) nonneg b negSucc c
additionZInjectiveSecondLemmaLemma zero zero zero pr = naughtE (negSuccInjective pr)
additionZInjectiveSecondLemmaLemma zero zero (succ c) pr = naughtE (negSuccInjective pr)
additionZInjectiveSecondLemmaLemma zero (succ b) c pr = exFalso (nonnegIsNotNegsucc pr)
additionZInjectiveSecondLemmaLemma (succ a) zero c pr = naughtE (canSubtractFromEqualityLeft {succ a} pr')
where
pr' : succ a +N zero succ a +N succ c
pr' rewrite addZeroRight (succ a) = negSuccInjective pr
additionZInjectiveSecondLemmaLemma (succ a) (succ b) c pr = naughtE r
where
pr' : negSucc (succ (a +N succ c)) +Z nonneg (succ a) nonneg b
pr' = equalityCommutative (moveNegsucc a (nonneg b) (negSucc (succ (a +N succ c))) pr)
pr'' : nonneg (succ a) nonneg b +Z nonneg (succ (succ (a +N succ c)))
pr'' = moveNegsucc (succ (a +N succ c)) (nonneg (succ a)) (nonneg b) pr'
pr''' : nonneg (succ a) nonneg (b +N succ (succ (a +N succ c)))
pr''' rewrite equalityCommutative (addingNonnegIsHom b (succ (succ (a +N succ c)))) = pr''
pr'''' : succ a b +N ((succ (succ a)) +N succ c)
pr'''' = nonnegInjective pr'''
q : succ a (b +N (succ (succ a))) +N succ c
q = identityOfIndiscernablesRight (succ a) (b +N ((succ (succ a)) +N succ c)) ((b +N (succ (succ a))) +N succ c) _≡_ pr'''' (equalityCommutative (additionNIsAssociative b (succ (succ a)) (succ c)))
moveSucc : (a b : ) a +N succ b succ a +N b
moveSucc a b rewrite additionNIsCommutative a (succ b) | additionNIsCommutative a b = refl
q' : succ a (succ b +N succ a) +N succ c
q' = identityOfIndiscernablesRight (succ a) ((b +N succ (succ a)) +N succ c) ((succ b +N succ a) +N succ c) _≡_ q (applyEquality (λ t t +N succ c) (moveSucc b (succ a)))
q'' : succ a (succ a +N succ b) +N succ c
q'' = identityOfIndiscernablesRight (succ a) ((succ b +N succ a) +N succ c) ((succ a +N succ b) +N succ c) _≡_ q' (applyEquality (λ t t +N succ c) (additionNIsCommutative (succ b) (succ a)))
q''' : succ a succ a +N (succ b +N succ c)
q''' rewrite equalityCommutative (additionNIsAssociative (succ a) (succ b) (succ c)) = q''
q'''' : succ a +N zero succ a +N (succ b +N succ c)
q'''' rewrite addZeroRight (succ a) = q'''
r : zero succ b +N succ c
r = canSubtractFromEqualityLeft {succ a} q''''
additionZInjectiveSecondLemma : (a : ) (b c : ) (negSucc a +Z b negSucc a +Z c) b c
additionZInjectiveSecondLemma zero (nonneg zero) (nonneg zero) pr = refl
additionZInjectiveSecondLemma zero (nonneg zero) (nonneg (succ c)) pr = exFalso (nonnegIsNotNegsucc (equalityCommutative pr))
additionZInjectiveSecondLemma zero (nonneg (succ b)) (nonneg zero) pr = exFalso (nonnegIsNotNegsucc pr)
additionZInjectiveSecondLemma zero (nonneg (succ b)) (nonneg (succ c)) pr rewrite additionZIsCommutative (negSucc zero) (nonneg b) | additionZIsCommutative (negSucc zero) (nonneg c) = applyEquality (λ t nonneg (succ t)) pr'
where
pr' : b c
pr' = nonnegInjective pr
additionZInjectiveSecondLemma zero (nonneg zero) (negSucc c) pr = naughtE pr'
where
pr' : zero succ c
pr' = negSuccInjective pr
additionZInjectiveSecondLemma zero (negSucc b) (nonneg zero) pr = naughtE (negSuccInjective (equalityCommutative pr))
additionZInjectiveSecondLemma zero (negSucc b) (nonneg (succ c)) pr = exFalso (nonnegIsNotNegsucc (equalityCommutative pr))
additionZInjectiveSecondLemma zero (negSucc b) (negSucc c) pr = applyEquality negSucc (succInjective pr')
where
pr' : succ b succ c
pr' = negSuccInjective pr
additionZInjectiveSecondLemma zero (nonneg (succ b)) (negSucc c) pr = exFalso (nonnegIsNotNegsucc pr)
additionZInjectiveSecondLemma (succ a) (nonneg zero) (nonneg zero) pr = refl
additionZInjectiveSecondLemma (succ a) (nonneg zero) (nonneg (succ c)) pr = exFalso (nonnegIsNotNegsucc pr'')
where
pr' : nonneg c negSucc (succ a) +Z nonneg (succ a)
pr' = moveNegsucc a (nonneg c) (negSucc (succ a)) (equalityCommutative pr)
pr'' : nonneg c negSucc zero
pr'' rewrite equalityCommutative (addToNegativeSelf (succ a)) = pr'
additionZInjectiveSecondLemma (succ a) (nonneg (succ b)) (nonneg zero) pr = exFalso (nonnegIsNotNegsucc pr'')
where
pr' : nonneg b negSucc (succ a) +Z nonneg (succ a)
pr' = moveNegsucc a (nonneg b) (negSucc (succ a)) pr
pr'' : nonneg b negSucc zero
pr'' rewrite equalityCommutative (addToNegativeSelf (succ a)) = pr'
additionZInjectiveSecondLemma (succ a) (nonneg (succ b)) (nonneg (succ c)) pr = applyEquality (λ t nonneg (succ t)) pr''
where
pr' : nonneg b nonneg c
pr' = additionZInjectiveSecondLemma a (nonneg b) (nonneg c) pr
pr'' : b c
pr'' = nonnegInjective pr'
additionZInjectiveSecondLemma (succ a) (nonneg zero) (negSucc c) pr = naughtE pr''
where
pr' : succ a +N zero succ a +N succ c
pr' rewrite addZeroRight a = negSuccInjective pr
pr'' : zero succ c
pr'' = canSubtractFromEqualityLeft {succ a} pr'
additionZInjectiveSecondLemma (succ a) (nonneg (succ b)) (negSucc c) pr = additionZInjectiveSecondLemmaLemma (succ a) (succ b) c pr
additionZInjectiveSecondLemma (succ a) (negSucc b) (nonneg c) pr = equalityCommutative pr'
where
pr' : nonneg c negSucc b
pr' = additionZInjectiveSecondLemmaLemma (succ a) c b (equalityCommutative pr)
additionZInjectiveSecondLemma (succ a) (negSucc b) (negSucc c) pr = applyEquality negSucc pr'''
where
pr' : succ a +N succ b succ a +N succ c
pr' = negSuccInjective pr
pr'' : succ b succ c
pr'' = canSubtractFromEqualityLeft {succ a} pr'
pr''' : b c
pr''' = succInjective pr''
additionZInjective : (a b c : ) (a +Z b a +Z c) b c
additionZInjective (nonneg a) b c pr = additionZInjectiveFirstLemma a b c pr
additionZInjective (negSucc a) b c pr = additionZInjectiveSecondLemma a b c pr
addNonnegSucc : (a : ) (b c : ) (a +Z nonneg b nonneg c) a +Z nonneg (succ b) nonneg (succ c)
addNonnegSucc (nonneg x) b c pr rewrite addingNonnegIsHom x (succ b) | addingNonnegIsHom x b = applyEquality nonneg g
where
p : x +N b c
p = nonnegInjective pr
g : x +N succ b succ c
g = identityOfIndiscernablesLeft (succ (x +N b)) (succ c) (x +N succ b) _≡_ (applyEquality succ p) (equalityCommutative (succExtracts x b))
addNonnegSucc (negSucc x) b c pr = negSuccPlusNonnegNonnegConverse x (succ b) (succ c) (applyEquality succ p')
where
p' : b c +N succ x
p' = negSuccPlusNonnegNonneg x b c pr
addNonnegSuccLemma : (a x d : ) (negSucc (succ a) +Z nonneg zero negSucc (succ x) +Z nonneg (succ d)) negSucc (succ a) +Z nonneg (succ zero) negSucc (succ x) +Z nonneg (succ (succ d))
addNonnegSuccLemma a x d pr = s
where
p' : negSucc (succ a) +Z nonneg (succ x) nonneg d
p' = equalityCommutative (moveNegsucc x (nonneg d) (negSucc (succ a)) (equalityCommutative pr))
p'' : nonneg (succ x) nonneg d +Z nonneg (succ (succ a))
p'' = moveNegsucc (succ a) (nonneg (succ x)) (nonneg d) p'
p''' : nonneg (succ x) nonneg (d +N succ (succ a))
p''' rewrite equalityCommutative (addingNonnegIsHom d (succ (succ a))) = p''
q : succ x d +N succ (succ a)
q = nonnegInjective p'''
q' : succ x succ (succ a) +N d
q' rewrite additionNIsCommutative (succ (succ a)) d = q
q'' : succ x succ d +N succ a
q'' rewrite additionNIsCommutative d (succ a) = q'
q''' : succ x succ a +N succ d
q''' rewrite additionNIsCommutative (succ a) (succ d) = q''
r : nonneg (succ x) nonneg (succ a +N succ d)
r = applyEquality nonneg q'''
r' : nonneg (succ x) nonneg (succ a) +Z nonneg (succ d)
r' = identityOfIndiscernablesRight (nonneg (succ x)) (nonneg (succ a +N succ d)) (nonneg (succ a) +Z nonneg (succ d)) _≡_ r (addingNonnegIsHom (succ a) (succ d))
r'' : nonneg (succ x) nonneg (succ d) +Z nonneg (succ a)
r'' rewrite additionZIsCommutative (nonneg (succ d)) (nonneg (succ a)) = r'
r''' : nonneg (succ d) negSucc a +Z nonneg (succ x)
r''' = equalityCommutative (moveNegsuccConverse a (nonneg (succ x)) (nonneg (succ d)) r'')
s : negSucc a negSucc x +Z nonneg (succ d)
s = equalityCommutative (moveNegsuccConverse x (nonneg (succ d)) (negSucc a) r''')
addNonnegSucc' : (a b : ) (c d : ) (a +Z nonneg c b +Z nonneg d) a +Z nonneg (succ c) b +Z nonneg (succ d)
addNonnegSucc' a (nonneg x) c d pr rewrite addingNonnegIsHom x (succ d) | addingNonnegIsHom x d | addNonnegSucc a c (x +N d) pr = applyEquality nonneg (equalityCommutative (succExtracts x d))
addNonnegSucc' (nonneg a) (negSucc x) c d pr = equalityCommutative (moveNegsuccConverse x (nonneg (succ d)) ((nonneg a) +Z nonneg (succ c)) (equalityCommutative q))
where
p : ((nonneg a) +Z nonneg c) +Z nonneg (succ x) nonneg d
p = equalityCommutative (moveNegsucc x (nonneg d) ((nonneg a) +Z nonneg c) (equalityCommutative pr))
p' : nonneg ((a +N c) +N succ x) nonneg d
p' = identityOfIndiscernablesLeft ((nonneg a +Z nonneg c) +Z nonneg (succ x)) (nonneg d) (nonneg ((a +N c) +N succ x)) _≡_ p g
where
g : (nonneg a +Z nonneg c) +Z nonneg (succ x) nonneg ((a +N c) +N succ x)
g rewrite addingNonnegIsHom a c | addingNonnegIsHom (a +N c) (succ x) = refl
p'' : (a +N c) +N succ x d
p'' = nonnegInjective p'
p''' : (a +N succ c) +N succ x succ d
p''' = identityOfIndiscernablesLeft (succ (a +N c) +N succ x) (succ d) ((a +N succ c) +N succ x) _≡_ (applyEquality succ p'') g
where
g : succ ((a +N c) +N succ x) (a +N succ c) +N succ x
g = applyEquality (λ i i +N succ x) {succ (a +N c)} {a +N succ c} (equalityCommutative (succExtracts a c))
q : (nonneg a +Z nonneg (succ c)) +Z nonneg (succ x) nonneg (succ d)
q rewrite (addingNonnegIsHom a (succ c)) | addingNonnegIsHom (a +N succ c) (succ x) = applyEquality nonneg p'''
addNonnegSucc' (negSucc a) (negSucc x) c d pr with (inspect (negSucc x +Z nonneg d))
addNonnegSucc' (negSucc a) (negSucc x) c d pr | (nonneg int) with p = identityOfIndiscernablesRight (negSucc a +Z nonneg (succ c)) (nonneg (succ int)) (negSucc x +Z nonneg (succ d)) _≡_ (addNonnegSucc (negSucc a) c int (transitivity pr p)) (equalityCommutative (addNonnegSucc (negSucc x) d int p))
addNonnegSucc' (negSucc zero) (negSucc zero) c d pr | negSucc int with p = additionZInjective (negSucc zero) (nonneg c) (nonneg d) pr
where
pr' : nonneg c (negSucc zero +Z nonneg d) +Z nonneg 1
pr' = moveNegsucc zero (nonneg c) (negSucc zero +Z nonneg d) pr
addNonnegSucc' (negSucc zero) (negSucc (succ x)) zero d pr | negSucc int with p = equalityCommutative (moveNegsuccConverse x (nonneg d) (nonneg zero) (equalityCommutative (applyEquality nonneg q')))
where
pr' : nonneg d (negSucc zero) +Z (nonneg (succ (succ x)))
pr' = moveNegsucc (succ x) (nonneg d) (negSucc zero) (equalityCommutative pr)
pr'' : nonneg (succ (succ x)) nonneg d +Z nonneg 1
pr'' = moveNegsucc (zero) (nonneg (succ (succ x))) (nonneg d) (equalityCommutative pr')
pr''' : nonneg (succ (succ x)) nonneg (d +N 1)
pr''' rewrite equalityCommutative (addingNonnegIsHom d 1) = pr''
pr'''' : succ (succ x) d +N 1
pr'''' = nonnegInjective pr'''
q : succ (succ x) succ d
q rewrite additionNIsCommutative 1 d = pr''''
q' : succ x d
q' = succInjective q
addNonnegSucc' (negSucc zero) (negSucc (succ x)) (succ c) zero pr | negSucc int with p = exFalso (nonnegIsNotNegsucc pr)
addNonnegSucc' (negSucc zero) (negSucc (succ x)) (succ c) (succ d) pr | negSucc int with p = addNonnegSucc' (nonneg zero) (negSucc x) c d pr
addNonnegSucc' (negSucc (succ a)) (negSucc zero) zero d pr | negSucc int with p = naughtE q'
where
pr' : negSucc (succ a) +Z nonneg 1 nonneg d
pr' = equalityCommutative (moveNegsucc zero (nonneg d) (negSucc (succ a)) (equalityCommutative pr))
pr'' : nonneg 1 nonneg d +Z nonneg (succ (succ a))
pr'' = moveNegsucc (succ a) (nonneg 1) (nonneg d) pr'
pr''' : nonneg 1 nonneg (d +N succ (succ a))
pr''' rewrite equalityCommutative (addingNonnegIsHom d (succ (succ a))) = pr''
q : 1 succ (succ a) +N d
q rewrite additionNIsCommutative (succ (succ a)) d = nonnegInjective pr'''
q' : 0 succ a +N d
q' = succInjective q
addNonnegSucc' (negSucc (succ a)) (negSucc zero) (succ c) zero pr | negSucc int with p = help
where
pr' : negSucc zero +Z nonneg (succ a) nonneg c
pr' = equalityCommutative (moveNegsucc a (nonneg c) (negSucc zero) pr)
pr'' : nonneg (succ a) nonneg c +Z nonneg 1
pr'' = moveNegsucc zero (nonneg (succ a)) (nonneg c) pr'
pr''' : nonneg (succ a) nonneg (c +N 1)
pr''' rewrite equalityCommutative (addingNonnegIsHom c 1) = pr''
q : succ a succ c
q rewrite additionNIsCommutative 1 c = nonnegInjective pr'''
q' : a c
q' = succInjective q
help : negSucc a +Z nonneg (succ c) nonneg zero
help rewrite q' = additiveInverseExists c
addNonnegSucc' (negSucc (succ a)) (negSucc zero) (succ c) (succ d) pr | negSucc int with p = moveNegsuccConverse a (nonneg (succ c)) (nonneg (succ d)) q'
where
pr' : nonneg c nonneg d +Z nonneg (succ a)
pr' = moveNegsucc a (nonneg c) (nonneg d) pr
pr'' : nonneg c nonneg (d +N succ a)
pr'' rewrite equalityCommutative (addingNonnegIsHom d (succ a)) = pr'
pr''' : c d +N succ a
pr''' = nonnegInjective pr''
pr'''' : succ c succ d +N succ a
pr'''' = applyEquality succ pr'''
q : nonneg (succ c) nonneg (succ d +N succ a)
q = applyEquality nonneg pr''''
q' : nonneg (succ c) nonneg (succ d) +Z nonneg (succ a)
q' rewrite addingNonnegIsHom (succ d) (succ a) = q
addNonnegSucc' (negSucc (succ a)) (negSucc (succ x)) zero zero pr | negSucc int with p = applyEquality negSucc (succInjective (negSuccInjective pr))
addNonnegSucc' (negSucc (succ a)) (negSucc (succ x)) zero (succ d) pr | negSucc int with p = addNonnegSuccLemma a x d pr
addNonnegSucc' (negSucc (succ a)) (negSucc (succ x)) (succ c) zero pr | negSucc int with p = equalityCommutative (addNonnegSuccLemma x a c (equalityCommutative pr))
addNonnegSucc' (negSucc (succ a)) (negSucc (succ x)) (succ c) (succ d) pr | negSucc int with p = addNonnegSucc' (negSucc a) (negSucc x) c d pr
subtractNonnegSucc : (a : ) (b c : ) (a +Z nonneg (succ b) nonneg (succ c)) a +Z nonneg b nonneg c
subtractNonnegSucc (nonneg x) b c pr rewrite addingNonnegIsHom x (succ b) | addingNonnegIsHom x b = applyEquality nonneg g
where
p : succ (x +N b) succ c
p rewrite succExtracts x b = nonnegInjective pr
g : x +N b c
g = succInjective p
subtractNonnegSucc (negSucc x) b c pr = negSuccPlusNonnegNonnegConverse x b c (succInjective p')
where
p' : succ b succ (c +N succ x)
p' = negSuccPlusNonnegNonneg x (succ b) (succ c) pr
addNegsuccThenNonneg : (a : ) (b c : ) (a +Z negSucc b) +Z nonneg (succ (c +N b)) a +Z nonneg c
addNegsuccThenNonneg (nonneg zero) zero c rewrite addZeroRight c = refl
addNegsuccThenNonneg (nonneg (succ x)) zero c rewrite addZeroRight c | addingNonnegIsHom x (succ c) = applyEquality nonneg (succExtracts x c)
addNegsuccThenNonneg (nonneg zero) (succ b) c = moveNegsuccConverse b (nonneg (c +N succ b)) (nonneg c) (equalityCommutative (addingNonnegIsHom c (succ b)))
addNegsuccThenNonneg (nonneg (succ x)) (succ b) c with addNegsuccThenNonneg (nonneg x) b c
... | prev = go
where
go : (nonneg x +Z negSucc b) +Z nonneg (succ (c +N succ b)) nonneg (succ (x +N c))
go rewrite addingNonnegIsHom x c | succExtracts c b = p
where
p : (nonneg x +Z negSucc b) +Z nonneg (succ (succ (c +N b))) nonneg (succ (x +N c))
p = addNonnegSucc (nonneg x +Z negSucc b) (succ (c +N b)) (x +N c) prev
addNegsuccThenNonneg (negSucc x) zero c rewrite addZeroRight c | sumOfNegsucc x 0 | addZeroRight x = refl
addNegsuccThenNonneg (negSucc x) (succ b) c with addNegsuccThenNonneg (negSucc x) b c
... | prev = go
where
p : nonneg c ((negSucc x +Z negSucc b) +Z nonneg (succ (c +N b))) +Z nonneg (succ x)
p = moveNegsucc x (nonneg c) ((negSucc x +Z negSucc b) +Z nonneg (succ (c +N b))) (equalityCommutative prev)
go : (negSucc x +Z negSucc (succ b)) +Z nonneg (succ (c +N succ b)) negSucc x +Z nonneg c
go rewrite sumOfNegsucc x (succ b) | succExtracts c b = identityOfIndiscernablesLeft ((negSucc x +Z negSucc b) +Z nonneg (succ (c +N b))) (negSucc x +Z nonneg c) (negSucc (x +N succ b) +Z nonneg (succ (c +N b))) _≡_ prev (applyEquality (λ t t +Z nonneg (succ (c +N b))) {negSucc x +Z negSucc b} {negSucc (x +N succ b)} (identityOfIndiscernablesRight (negSucc x +Z negSucc b) (negSucc (succ (x +N b))) (negSucc (x +N succ b)) _≡_ (sumOfNegsucc x b) (applyEquality negSucc (equalityCommutative (succExtracts x b)))))
addNonnegThenNonneg : (a : ) (b c : ) (a +Z nonneg b) +Z nonneg c a +Z nonneg (b +N c)
addNonnegThenNonneg (nonneg x) b c rewrite addingNonnegIsHom x b | addingNonnegIsHom x (b +N c) | addingNonnegIsHom (x +N b) c = applyEquality nonneg (additionNIsAssociative x b c)
addNonnegThenNonneg (negSucc x) b zero rewrite addZeroRight b | additionZeroDoesNothing (negSucc x +Z nonneg b) = refl
addNonnegThenNonneg (negSucc x) b (succ c) with addNonnegThenNonneg (negSucc x) b c
... | prev = prev'
where
prev' : ((negSucc x +Z nonneg b) +Z nonneg (succ c)) negSucc x +Z nonneg (b +N succ c)
prev' rewrite addNonnegSucc' (negSucc x +Z nonneg b) (negSucc x) c (b +N c) prev | succExtracts b c = refl
additionZIsAssociativeFirstAndSecondArgNonneg : (a b : ) (c : ) (nonneg a +Z (nonneg b +Z c)) ((nonneg a) +Z nonneg b) +Z c
additionZIsAssociativeFirstAndSecondArgNonneg zero b c = refl
additionZIsAssociativeFirstAndSecondArgNonneg (succ a) b (nonneg zero) = i
where
h : nonneg (succ a) +Z (nonneg b +Z nonneg zero) nonneg (succ a) +Z nonneg b
h = identityOfIndiscernablesLeft (nonneg (succ a) +Z nonneg b) (nonneg (succ a) +Z nonneg b) (nonneg (succ a) +Z (nonneg b +Z nonneg zero)) _≡_ refl (applyEquality (λ r nonneg (succ a) +Z r) {nonneg b} {nonneg b +Z nonneg zero} (equalityCommutative (additionZeroDoesNothing (nonneg b))))
i : nonneg (succ a) +Z (nonneg b +Z nonneg zero) (nonneg (succ a) +Z nonneg b) +Z nonneg zero
i = identityOfIndiscernablesRight (nonneg (succ a) +Z (nonneg b +Z nonneg zero)) (nonneg (succ a) +Z nonneg b) ((nonneg (succ a) +Z nonneg b) +Z nonneg zero) _≡_ h (equalityCommutative (additionZeroDoesNothing (nonneg (succ a) +Z nonneg b)))
additionZIsAssociativeFirstAndSecondArgNonneg (succ x) zero (nonneg (succ c)) = applyEquality nonneg (applyEquality succ (applyEquality (λ n n +N succ c) {x} {x +N zero} (equalityCommutative (addZeroRight x))))
additionZIsAssociativeFirstAndSecondArgNonneg (succ x) (succ b) (nonneg (succ c)) = applyEquality nonneg (applyEquality succ i)
where
h : x +N (succ b +N succ c) (x +N succ b) +N succ c
h = equalityCommutative (additionNIsAssociative x (succ b) (succ c))
i : x +N succ (b +N succ c) (x +N succ b) +N succ c
i = identityOfIndiscernablesLeft (x +N (succ b +N succ c)) ((x +N succ b) +N succ c) (x +N succ (b +N succ c)) _≡_ h refl
additionZIsAssociativeFirstAndSecondArgNonneg (succ x) zero (negSucc c) rewrite addZeroRight x = refl
additionZIsAssociativeFirstAndSecondArgNonneg (succ x) (succ b) (negSucc zero) rewrite additionNIsCommutative x b | additionNIsCommutative x (succ b) = refl
additionZIsAssociativeFirstAndSecondArgNonneg (succ x) (succ b) (negSucc (succ c)) rewrite additionNIsCommutative x (succ b) | additionNIsCommutative b x = additionZIsAssociativeFirstAndSecondArgNonneg (succ x) b (negSucc c)
additionZIsAssociativeFirstArgNonneg : (a : ) (b c : ) (nonneg a +Z (b +Z c) ((nonneg a) +Z b) +Z c)
additionZIsAssociativeFirstArgNonneg zero (nonneg b) c = additionZIsAssociativeFirstAndSecondArgNonneg 0 b c
additionZIsAssociativeFirstArgNonneg 0 (negSucc b) c = refl
additionZIsAssociativeFirstArgNonneg (succ a) (nonneg b) c = additionZIsAssociativeFirstAndSecondArgNonneg (succ a) b c
additionZIsAssociativeFirstArgNonneg (succ x) (negSucc zero) (nonneg 0) rewrite addingNonnegIsHom x zero | addZeroRight x = refl
additionZIsAssociativeFirstArgNonneg (succ x) (negSucc zero) (nonneg (succ c)) = i
where
h : nonneg (succ (x +N c)) nonneg (x +N succ c)
s : nonneg (x +N succ c) nonneg (succ c +N x)
t : nonneg (x +N succ c) nonneg (succ c) +Z nonneg x
i : nonneg (succ (x +N c)) nonneg x +Z nonneg (succ c)
h = applyEquality nonneg (equalityCommutative (succExtracts x c))
s = applyEquality nonneg (additionNIsCommutative x (succ c))
t = transitivity s refl
i = transitivity h (equalityCommutative (addingNonnegIsHom x (succ c)))
additionZIsAssociativeFirstArgNonneg (succ x) (negSucc (succ b)) (nonneg 0) rewrite additionZIsCommutative (nonneg x +Z negSucc b) (nonneg zero) = refl
additionZIsAssociativeFirstArgNonneg (succ x) (negSucc (succ b)) (nonneg (succ c)) = p
where
p''' : nonneg x +Z (negSucc b +Z nonneg c) (nonneg x +Z negSucc b) +Z nonneg c
p''' = additionZIsAssociativeFirstArgNonneg x (negSucc b) (nonneg c)
p'' : (negSucc b +Z nonneg c) +Z nonneg x (nonneg x +Z negSucc b) +Z nonneg c
p'' = identityOfIndiscernablesLeft (nonneg x +Z (negSucc b +Z nonneg c)) ((nonneg x +Z negSucc b) +Z nonneg c) ((negSucc b +Z nonneg c) +Z nonneg x) _≡_ p''' (additionZIsCommutative (nonneg x) (negSucc b +Z nonneg c))
p' : (negSucc b +Z nonneg c) +Z nonneg (succ x) (nonneg x +Z negSucc b) +Z nonneg (succ c)
p' = addNonnegSucc' (negSucc b +Z nonneg c) (nonneg x +Z negSucc b) x c p''
p : nonneg (succ x) +Z (negSucc b +Z nonneg c) (nonneg x +Z negSucc b) +Z nonneg (succ c)
p = identityOfIndiscernablesLeft ((negSucc b +Z nonneg c) +Z nonneg (succ x)) ((nonneg x +Z negSucc b) +Z nonneg (succ c)) (nonneg (succ x) +Z (negSucc b +Z nonneg c)) _≡_ p' (additionZIsCommutative (negSucc b +Z nonneg c) (nonneg (succ x)))
additionZIsAssociativeFirstArgNonneg (succ x) (negSucc zero) (negSucc c) = refl
additionZIsAssociativeFirstArgNonneg (succ a) (negSucc (succ b)) (negSucc zero) rewrite additionNIsCommutative b 1 | additionZIsCommutative (nonneg a +Z negSucc b) (negSucc 0) = equalityCommutative (moveNegsuccConverse 0 (nonneg a +Z negSucc b) (nonneg a +Z negSucc (succ b)) q'')
where
q : nonneg 1 +Z (nonneg a +Z negSucc (succ b)) (nonneg 1 +Z nonneg a) +Z negSucc (succ b)
q = additionZIsAssociativeFirstAndSecondArgNonneg 1 a (negSucc (succ b))
q' : (nonneg 1 +Z nonneg a) +Z negSucc (succ b) (nonneg a +Z negSucc (succ b)) +Z nonneg 1
q' rewrite additionZIsCommutative (nonneg a +Z negSucc (succ b)) (nonneg 1) = equalityCommutative q
q'' : nonneg (succ a) +Z negSucc (succ b) (nonneg a +Z negSucc (succ b)) +Z nonneg 1
q'' rewrite addingNonnegIsHom 1 a = q'
additionZIsAssociativeFirstArgNonneg (succ a) (negSucc (succ b)) (negSucc (succ c)) = ans
where
ans : nonneg a +Z negSucc (b +N succ (succ c)) (nonneg a +Z negSucc b) +Z negSucc (succ c)
p : nonneg a +Z (negSucc b +Z negSucc (succ c)) (nonneg a +Z negSucc b) +Z negSucc (succ c)
p = additionZIsAssociativeFirstArgNonneg a (negSucc b) (negSucc (succ c))
p' : negSucc (b +N succ (succ c)) negSucc b +Z negSucc (succ c)
p' rewrite additionZIsCommutative (negSucc b) (negSucc (succ c)) | additionNIsCommutative b (succ (succ c)) | additionNIsCommutative c (succ b) | additionNIsCommutative c b = refl
ans rewrite p' = p
additionZIsAssociative : (a b c : ) (a +Z (b +Z c)) (a +Z b) +Z c
additionZIsAssociative (nonneg a) b c = additionZIsAssociativeFirstArgNonneg a b c
additionZIsAssociative (negSucc a) (nonneg b) c rewrite additionZIsCommutative (negSucc a) (nonneg b) | additionZIsCommutative (negSucc a) (nonneg b +Z c) = p''
where
p : (nonneg b +Z c) +Z negSucc a nonneg b +Z (c +Z negSucc a)
p = equalityCommutative (additionZIsAssociativeFirstArgNonneg b c (negSucc a))
p' : (nonneg b +Z c) +Z negSucc a nonneg b +Z (negSucc a +Z c)
p' rewrite additionZIsCommutative (negSucc a) c = p
p'' : (nonneg b +Z c) +Z negSucc a (nonneg b +Z negSucc a) +Z c
p'' rewrite equalityCommutative (additionZIsAssociativeFirstArgNonneg b (negSucc a) c) = p'
additionZIsAssociative (negSucc a) (negSucc b) (nonneg c) rewrite additionZIsCommutative (negSucc a +Z negSucc b) (nonneg c) | additionZIsCommutative (negSucc b) (nonneg c) | additionZIsCommutative (negSucc a) (nonneg c +Z negSucc b) = p'
where
p : (nonneg c +Z negSucc b) +Z negSucc a nonneg c +Z (negSucc b +Z negSucc a)
p = equalityCommutative (additionZIsAssociativeFirstArgNonneg c (negSucc b) (negSucc a))
p' : (nonneg c +Z negSucc b) +Z negSucc a nonneg c +Z (negSucc a +Z negSucc b)
p' rewrite additionZIsCommutative (negSucc a) (negSucc b) = p
additionZIsAssociative (negSucc zero) (negSucc zero) (negSucc c) = refl
additionZIsAssociative (negSucc zero) (negSucc (succ b)) (negSucc c) = refl
additionZIsAssociative (negSucc (succ a)) (negSucc zero) (negSucc c) rewrite additionNIsCommutative a 1 | additionNIsCommutative a (succ (succ c)) | additionNIsCommutative a (succ c) = refl
additionZIsAssociative (negSucc (succ a)) (negSucc (succ b)) (negSucc zero) rewrite additionNIsCommutative b 1 | additionNIsCommutative (succ ((a +N succ (succ b)))) 1 = applyEquality (λ t negSucc (succ t)) (p (succ (succ b)))
where
p : (r : ) a +N succ r succ (a +N r)
p r rewrite additionNIsCommutative a (succ r) | additionNIsCommutative r a = refl
additionZIsAssociative (negSucc (succ a)) (negSucc (succ b)) (negSucc (succ c)) = applyEquality (λ t negSucc (succ t)) (equalityCommutative (additionNIsAssociative a (succ (succ b)) (succ (succ c))))
lessNegsuccNonneg : {a b : } (negSucc a <Z nonneg b)
lessNegsuccNonneg {a} {b} = record { x = x ; proof = pr }
where
x :
x = a +N b
pr' : nonneg (succ (a +N b)) nonneg b +Z nonneg (succ a)
pr' rewrite addingNonnegIsHom b (succ a) | additionNIsCommutative b (succ a) = refl
pr : nonneg (succ x) +Z negSucc a nonneg b
pr rewrite additionZIsCommutative (nonneg (succ x)) (negSucc a) = moveNegsuccConverse a (nonneg (succ (a +N b))) (nonneg b) pr'
moveNegsuccTimes : (a b : ) (c : ) (a *Z (negSucc c)) b b +Z (a *Z nonneg (succ c)) nonneg 0
moveNegsuccTimes (nonneg zero) (nonneg b) c pr rewrite multiplyWithZero' (negSucc c) | additionZIsCommutative (nonneg b) (nonneg 0) = equalityCommutative pr
moveNegsuccTimes (nonneg (succ a)) (nonneg b) zero ()
moveNegsuccTimes (nonneg (succ a)) (nonneg b) (succ c) pr = {!!}
moveNegsuccTimes (negSucc a) (nonneg b) c pr = {!!}
moveNegsuccTimes (nonneg 0) (negSucc b) c pr = {!!}
moveNegsuccTimes (nonneg (succ a)) (negSucc b) c pr = {!!}
moveNegsuccTimes (negSucc a) (negSucc b) c pr = {!!}
multiplicationZIsAssociative : (a b c : ) (a *Z (b *Z c)) (a *Z b) *Z c
multiplicationZIsAssociative (nonneg x) (nonneg x₁) (nonneg x₂) = applyEquality nonneg (multiplicationNIsAssociative x x₁ x₂)
multiplicationZIsAssociative (nonneg x) (nonneg zero) (negSucc zero) = p
where
a : x *N zero zero
a = productZeroIsZeroRight x
q : nonneg zero nonneg zero *Z negSucc zero
q = refl
r : nonneg (x *N zero) nonneg zero *Z negSucc zero
r = identityOfIndiscernablesLeft (nonneg zero) (nonneg zero *Z negSucc zero) (nonneg (x *N zero)) _≡_ q (applyEquality nonneg (equalityCommutative a))
r' : nonneg zero *Z negSucc zero nonneg (x *N zero) *Z negSucc zero
r' = applyEquality (λ n n *Z negSucc zero) (applyEquality nonneg (equalityCommutative a))
p : nonneg (x *N zero) nonneg (x *N zero) *Z negSucc zero
p = identityOfIndiscernablesRight (nonneg (x *N zero)) (nonneg zero *Z negSucc zero) (nonneg (x *N zero) *Z negSucc zero) _≡_ r r'
multiplicationZIsAssociative (nonneg zero) (nonneg (succ y)) (negSucc zero) = zeroMultInZLeft (negSucc y)
multiplicationZIsAssociative (nonneg zero) (nonneg (succ y)) (negSucc (succ z)) = zeroMultInZLeft (negSucc y +Z nonneg (succ y) *Z negSucc z)
multiplicationZIsAssociative (nonneg (succ x)) (nonneg (succ y)) (negSucc zero) = {!!}
-- moveNegsucc : (a : ) (b c : ) → (negSucc a) +Z b ≡ c → b ≡ c +Z (nonneg (succ a))
multiplicationZIsAssociative (nonneg (succ x)) (nonneg (succ y)) (negSucc (succ z)) = {!!}
multiplicationZIsAssociative (nonneg x) (negSucc b) (nonneg c) = {!!}
multiplicationZIsAssociative (nonneg x) (negSucc b) (negSucc c) = {!!}
multiplicationZIsAssociative (negSucc x) (nonneg b) (nonneg c) = {!!}
multiplicationZIsAssociative (negSucc x) (nonneg b) (negSucc c) = {!!}
multiplicationZIsAssociative (negSucc x) (negSucc b) (nonneg c) = {!!}
multiplicationZIsAssociative (negSucc x) (negSucc b) (negSucc c) = {!!}
multiplicationZIsAssociative (nonneg x) (nonneg zero) (negSucc (succ z)) = {!!}
zeroIsAdditiveIdentityRightZ : (a : ) (a +Z nonneg zero) a
zeroIsAdditiveIdentityRightZ (nonneg x) = identityOfIndiscernablesRight (nonneg x +Z nonneg zero) (nonneg (x +N zero)) (nonneg x) _≡_ (addingNonnegIsHom x zero) ((applyEquality nonneg (addZeroRight x)))
zeroIsAdditiveIdentityRightZ (negSucc a) = refl
additiveInverseZ : (a : )
additiveInverseZ (nonneg zero) = nonneg zero
additiveInverseZ (nonneg (succ x)) = negSucc x
additiveInverseZ (negSucc a) = nonneg (succ a)
addInverseLeftZ : (a : ) (additiveInverseZ a +Z a nonneg zero)
addInverseLeftZ (nonneg zero) = refl
addInverseLeftZ (nonneg (succ zero)) = refl
addInverseLeftZ (nonneg (succ (succ x))) = addInverseLeftZ (nonneg (succ x))
addInverseLeftZ (negSucc zero) = refl
addInverseLeftZ (negSucc (succ a)) = addInverseLeftZ (negSucc a)
addInverseRightZ : (a : ) (a +Z additiveInverseZ a nonneg zero)
addInverseRightZ (nonneg zero) = refl
addInverseRightZ (nonneg (succ zero)) = refl
addInverseRightZ (nonneg (succ (succ x))) = addInverseRightZ (nonneg (succ x))
addInverseRightZ (negSucc zero) = refl
addInverseRightZ (negSucc (succ a)) = addInverseRightZ (negSucc a)
addZDistributiveMulZ : (a b c : ) (a *Z (b +Z c)) (a *Z b +Z a *Z c)
addZDistributiveMulZ (nonneg a) (nonneg b) (nonneg c) rewrite addingNonnegIsHom b c | multiplyingNonnegIsHom a (b +N c) | addingNonnegIsHom (a *N b) (a *N c) = applyEquality nonneg (productDistributes a b c)
addZDistributiveMulZ (nonneg zero) (nonneg zero) (negSucc zero) = refl
addZDistributiveMulZ (nonneg zero) (nonneg zero) (negSucc (succ c)) = refl
addZDistributiveMulZ (nonneg zero) (nonneg (succ b)) (negSucc zero) = refl
addZDistributiveMulZ (nonneg zero) (nonneg (succ b)) (negSucc (succ c)) rewrite multiplicationZIsCommutative (nonneg zero) (nonneg b +Z negSucc c) = multiplyWithZero (nonneg b +Z negSucc c)
addZDistributiveMulZ (nonneg (succ a)) (nonneg zero) (negSucc zero) rewrite multiplicationNIsCommutative a zero = refl
addZDistributiveMulZ (nonneg (succ a)) (nonneg zero) (negSucc (succ c)) rewrite multiplicationNIsCommutative a zero = refl
addZDistributiveMulZ (nonneg (succ a)) (nonneg (succ b)) (negSucc zero) = identityOfIndiscernablesRight (nonneg (b +N (a *N b))) (negSucc a +Z nonneg (succ (b +N a *N succ b))) (nonneg (succ (b +N a *N succ b)) +Z negSucc a) _≡_ p (additionZIsCommutative (negSucc a) (nonneg (succ (b +N a *N succ b))))
where
p : nonneg (b +N (a *N b)) (negSucc a) +Z (nonneg (succ (b +N a *N succ b)))
q : nonneg (succ (b +N a *N succ b)) nonneg (b +N a *N b) +Z nonneg (succ a)
r : succ a *N succ b succ a +N (succ a *N b)
r' : succ a *N succ b succ a +N (b *N succ a)
r' rewrite multiplicationNIsCommutative (succ a) (succ b) = refl
r rewrite multiplicationNIsCommutative (succ a) b = r'
p = equalityCommutative (moveNegsuccConverse a (nonneg (succ (b +N a *N succ b))) (nonneg (b +N a *N b)) q)
q rewrite addingNonnegIsHom (b +N a *N b) (succ a) | additionNIsCommutative (b +N a *N b) (succ a) = applyEquality (λ t nonneg t) r
addZDistributiveMulZ (nonneg (succ a)) (nonneg (succ b)) (negSucc (succ c)) = {!!}
addZDistributiveMulZ (nonneg zero) (negSucc b) (nonneg c) rewrite additionZIsCommutative (nonneg zero *Z negSucc b) (nonneg zero) | multiplyWithZero' (negSucc b) | multiplyWithZero' (negSucc b +Z nonneg c) = refl
addZDistributiveMulZ (nonneg (succ a)) (negSucc b) (nonneg c) = {!!}
addZDistributiveMulZ (nonneg zero) (negSucc b) (negSucc c) rewrite multiplyWithZero' (negSucc b +Z negSucc c) | multiplyWithZero' (negSucc b) | multiplyWithZero' (negSucc c) = refl
addZDistributiveMulZ (nonneg (succ a)) (negSucc zero) (negSucc c) = refl
addZDistributiveMulZ (nonneg (succ a)) (negSucc (succ b)) (negSucc zero) rewrite additionNIsCommutative b 1 | additionZIsCommutative (negSucc a +Z nonneg (succ a) *Z negSucc b) (negSucc a) = refl
addZDistributiveMulZ (nonneg (succ a)) (negSucc (succ b)) (negSucc (succ c)) = {!!}
addZDistributiveMulZ (negSucc a) b c = {!!}
zGroup : group {_} {}
group._·_ zGroup = _+Z_
group.identity zGroup = nonneg zero
group.inv zGroup = additiveInverseZ
group.multAssoc zGroup {a} {b} {c} = additionZIsAssociative a b c
group.multIdentLeft zGroup {a} = refl
group.multIdentRight zGroup {a} = zeroIsAdditiveIdentityRightZ a
group.invLeft zGroup {a} = addInverseLeftZ a
group.invRight zGroup {a} = addInverseRightZ a
zRing : ring {_} {}
ring.additiveGroup zRing = zGroup
ring._*_ zRing = _*Z_
ring.multIdent zRing = nonneg (succ zero)
ring.groupIsAbelian zRing {a} {b} = additionZIsCommutative a b
ring.multAssoc zRing {a} {b} {c} = multiplicationZIsAssociative a b c
ring.multCommutative zRing {a} {b} = multiplicationZIsCommutative a b
ring.multDistributes zRing {a} {b} {c} = addZDistributiveMulZ a b c
ring.multIdentIsIdent zRing {a} = multiplicativeIdentityOneZLeft a
-- TODO: ordered ring axioms

275
Vectors.agda Normal file
View File

@@ -0,0 +1,275 @@
{-# OPTIONS --warning=error --safe #-}
open import LogicalFormulae
open import Naturals
open import Functions
data Vec {a : _} (X : Set a) : -> Set a where
[] : Vec X zero
_,-_ : {n : } -> X -> Vec X n -> Vec X (succ n)
vecLen : {a : _} {X : Set a} {n : } Vec X n
vecLen {a} {X} {n} v = n
vecHead : {a : _} {X : Set a} {n : } Vec X (succ n) X
vecHead (x ,- xs) = x
vecTail : {a : _} {X : Set a} {n : } Vec X (succ n) Vec X n
vecTail (x ,- xs) = xs
vecIsHeadPlusTail : {a : _} {X : Set a} {n : } (xs : Vec X (succ n)) (vecHead xs ,- vecTail xs) xs
vecIsHeadPlusTail (x ,- xs) = refl
_+V_ : {a : _} {X : Set a} {m n : } Vec X m Vec X n Vec X (m +N n)
[] +V ys = ys
(x ,- xs) +V ys = (x ,- (xs +V ys))
vecIndex : {a : _} {X : Set a} {m : } Vec X m (i : ) (i <N m) X
vecIndex [] zero (le x ())
vecIndex (x ,- vec) zero i<m = x
vecIndex [] (succ i) (le x ())
vecIndex (x ,- vec) (succ i) i<m = vecIndex vec i (canRemoveSuccFrom<N i<m)
vecLast : {a : _} {X : Set a} {m : } Vec X m (0 <N m) X
vecLast {a} {X} {zero} v ()
vecLast {a} {X} {succ zero} (x ,- []) _ = x
vecLast {a} {X} {succ (succ m)} (x ,- v) _ = vecLast v (succIsPositive m)
vecAppend : {a : _} {X : Set a} {m : } Vec X m (x : X) Vec X (succ m)
vecAppend {a} {X} {zero} [] x = x ,- []
vecAppend {a} {X} {succ m} (y ,- v) x = y ,- vecAppend v x
vecAppendIsNowLast : {a : _} {X : Set a} {m : } (x : X) (v : Vec X m) (vecLast (vecAppend v x) (succIsPositive m)) x
vecAppendIsNowLast {a} {X} {zero} x [] = refl
vecAppendIsNowLast {a} {X} {succ m} x (y ,- v) = vecAppendIsNowLast x v
vecRev : {a : _} {X : Set a} {m : } Vec X m Vec X m
vecRev {a} {X} {zero} [] = []
vecRev {a} {X} {succ m} (x ,- v) = vecAppend (vecRev v) x
vecMoveAppend : {a : _} {X : Set a} {m : } (x : X) (v : Vec X m) vecRev (vecAppend v x) x ,- vecRev v
vecMoveAppend {a} {X} {.0} x [] = refl
vecMoveAppend {a} {X} {.(succ _)} x (y ,- v) rewrite vecMoveAppend x v = refl
revRevIsId : {a : _} {X : Set a} {m : } (v : Vec X m) (vecRev (vecRev v)) v
revRevIsId {a} {X} {zero} [] = refl
revRevIsId {a} {X} {succ m} (x ,- v) rewrite vecMoveAppend x (vecRev v) = applyEquality (λ i x ,- i) (revRevIsId v)
record vecContains {a : _} {X : Set a} {m : } (vec : Vec X m) (x : X) : Set a where
field
index :
index<m : index <N m
isHere : vecIndex vec index index<m x
vecSolelyContains : {a : _} {X : Set a} {m : X} (n : X) (vecContains (m ,- []) n) m n
vecSolelyContains {m} n record { index = zero ; index<m = _ ; isHere = isHere } = isHere
vecSolelyContains {m} n record { index = (succ index) ; index<m = (le x proof) ; isHere = isHere } = exFalso (f {x} {index} proof)
where
f : {x index : } succ (x +N succ index) 1 False
f {x} {index} pr rewrite additionNIsCommutative x (succ index) = naughtE (equalityCommutative (succInjective pr))
vecChop : {a : _} {X : Set a} (m : ) {n : } Vec X (m +N n) Vec X m && Vec X n
_&&_.fst (vecChop zero xs) = []
_&&_.fst (vecChop (succ m) (x ,- xs)) = (x ,- _&&_.fst (vecChop m xs))
_&&_.snd (vecChop zero xs) = xs
_&&_.snd (vecChop (succ m) (x ,- xs)) = _&&_.snd (vecChop m xs)
{-
vecIsChopThenAppend : {a : _} {X : Set a} {m n : } (xs : Vec X m) (ys : Vec X n) → vecChop m (xs +V ys) ≡ record { fst = xs ; snd = ys }
vecIsChopThenAppend {a} {X} {m} xs ys with vecChop m (xs +V ys)
vecIsChopThenAppend {a} {X} {zero} [] ys | record { fst = [] ; snd = snd } = {!!}
vecIsChopThenAppend {a} {X} {succ m} xs ys | bl = {!!}
-}
vecMap : {a b : _} {X : Set a} {Y : Set b} (X Y) {n : } Vec X n Vec Y n
vecMap f [] = []
vecMap f (x ,- vec) = f x ,- (vecMap f vec)
vecMapCommutesWithAppend : {a b : _} {X : Set a} {Y : Set b} (f : X Y) {n : } (x : X) (v : Vec X n) (vecMap f (vecAppend v x) vecAppend (vecMap f v) (f x))
vecMapCommutesWithAppend f {.0} x [] = refl
vecMapCommutesWithAppend f {.(succ _)} x (y ,- v) = applyEquality (λ i (f y ,- i)) (vecMapCommutesWithAppend f x v)
vecMapCommutesWithRev : {a b : _} {X : Set a} {Y : Set b} (f : X Y) {n : } (v : Vec X n) (vecMap f (vecRev v) vecRev (vecMap f v))
vecMapCommutesWithRev {a} {b} {X} {Y} f {.0} [] = refl
vecMapCommutesWithRev {a} {b} {X} {Y} f {.(succ _)} (x ,- v) rewrite vecMapCommutesWithAppend f x (vecRev v) = applyEquality (λ i vecAppend i (f x)) (vecMapCommutesWithRev f v)
vecIndex0AndAppend : {a : _} {X : Set a} {n : } (v : Vec X n) (0<n : 0 <N n) (x : X) vecIndex (vecAppend v x) 0 (succIsPositive n) vecIndex v 0 0<n
vecIndex0AndAppend [] () x
vecIndex0AndAppend (a ,- v) 0<n x = refl
vecIndexMAndAppend : {a : _} {X : Set a} {n : } (v : Vec X n) (x : X) (m : ) (m<n : m <N n) (pr : m <N succ n) vecIndex (vecAppend v x) m pr vecIndex v m m<n
vecIndexMAndAppend {n = n} v x zero m<n pr rewrite <NRefl pr (succIsPositive n) = vecIndex0AndAppend v m<n x
vecIndexMAndAppend [] x (succ m) ()
vecIndexMAndAppend {n = n} (y ,- v) x (succ m) m<n pr = vecIndexMAndAppend v x m (canRemoveSuccFrom<N m<n) (canRemoveSuccFrom<N pr)
vecMapCompose : {a b c : _} {X : Set a} {Y : Set b} {Z : Set c} (f : X Y) (g : Y Z) {n : } (v : Vec X n) (vecMap g (vecMap f v)) vecMap (λ i g (f i)) v
vecMapCompose f g [] = refl
vecMapCompose f g (x ,- v) = applyEquality (λ w (g (f x) ,- w)) (vecMapCompose f g v)
vecMapIdFact : {a : _} {X : Set a} {f : X X} (feq : (x : X) f x x) {n : } (xs : Vec X n) vecMap f xs xs
vecMapIdFact feq [] = refl
vecMapIdFact feq (x ,- xs) rewrite (feq x) | vecMapIdFact feq xs = refl
vecMapCompositionFact : {a b c : _} {X : Set a} {Y : Set b} {Z : Set c} {f : Y Z} {g : X Y} {h : X Z} (heq : (x : X) f (g x) h x) {n : } (xs : Vec X n) vecMap f (vecMap g xs) vecMap h xs
vecMapCompositionFact heq [] = refl
vecMapCompositionFact {a} {b} {c} {X} {Y} {Z} {f} {g} {h} heq (x ,- xs) rewrite heq x | vecMapCompositionFact {a} {b} {c} {X} {Y} {Z} {f} {g} {h} heq xs = refl
vecMapIsNatural : {a b : _} {X : Set a} {Y : Set b} (f : X Y) {m n : } (xs : Vec X m) (xs' : Vec X n) vecMap f (xs +V xs') (vecMap f xs +V vecMap f xs')
vecMapIsNatural f [] xs' = refl
vecMapIsNatural f (x ,- xs) xs' rewrite vecMapIsNatural f xs xs' = refl
vecPure : {a : _} {X : Set a} X {n : } Vec X n
vecPure x {zero} = []
vecPure x {succ n} = x ,- vecPure x {n}
_$V_ : {a b : _} {X : Set a} {Y : Set b} {n : } Vec (X Y) n Vec X n Vec Y n
[] $V [] = []
(f ,- fs) $V (x ,- xs) = f x ,- (fs $V xs)
infixl 3 _$V_
vec : {a b : _} {X : Set a} {Y : Set b} (X Y) {n : } Vec X n Vec Y n
vec f xs = (vecPure f) $V xs
vecZip : {a b : _} {X : Set a} {Y : Set b} {n : } Vec X n Vec Y n Vec (X && Y) n
vecZip [] [] = []
vecZip (x ,- xs) (x₁ ,- ys) = record { fst = x ; snd = x₁ } ,- vecZip xs ys
vecIdentity : {a : _} {X : Set a} {f : X X} (feq : (x : X) f x x) {n : } (xs : Vec X n) (vecPure f $V xs) xs
vecIdentity feq [] = refl
vecIdentity feq (x ,- xs) rewrite feq x | vecIdentity feq xs = refl
vecHom : {a b : _} {X : Set a} {Y : Set b} (f : X Y) (x : X) {n : } (vecPure f $V vecPure x) vecPure (f x) {n}
vecHom f x {zero} = refl
vecHom f x {succ n} rewrite vecHom f x {n} = refl
vecInterchange : {a b : _} {X : Set a} {Y : Set b} {n : } (fs : Vec (X Y) n) (x : X) (fs $V vecPure x) (vecPure (λ f f x) $V fs)
vecInterchange [] x = refl
vecInterchange (f ,- fs) x rewrite vecInterchange fs x = refl
vecComposition : {a b c : _} {X : Set a} {Y : Set b} {Z : Set c} {n : } (fs : Vec (Y Z) n) (gs : Vec (X Y) n) (xs : Vec X n) (vecPure (λ i j x i (j x)) $V fs $V gs $V xs) (fs $V (gs $V xs))
vecComposition [] [] [] = refl
vecComposition (f ,- fs) (g ,- gs) (x ,- xs) rewrite vecComposition fs gs xs = refl
------------
data _<=_ : Set where
oz : zero <= zero
os : {n m : } n <= m succ n <= succ m
o' : {n m : } n <= m n <= succ m
all0<=4 : Vec (0 <= 4) 1
all0<=4 = o' (o' (o' (o' oz))) ,- []
all1<=4 : Vec (1 <= 4) 4
all1<=4 = (o' ((o' (o' (os oz))))) ,- ((o' (o' (os (o' oz)))) ,- ((o' (os (o' (o' oz)))) ,- ((os (o' (o' (o' oz)))) ,- [])))
all2<=4 : Vec (2 <= 4) 6
all2<=4 = os (os (o' (o' oz))) ,- (os (o' (os (o' oz))) ,- (os (o' (o' (os oz))) ,- (o' (o' (os (os oz))) ,- (o' (os (o' (os oz))) ,- (o' (os (os (o' oz))) ,- [])))))
all3<=4 : Vec (3 <= 4) 4
all3<=4 = os (os (os (o' oz))) ,- (os (os (o' (os oz))) ,- (os (o' (os (os oz))) ,- (o' (os (os (os oz))) ,- [])))
all4<=4 : Vec (4 <= 4) 1
all4<=4 = os (os (os (os oz))) ,- []
<=Is≤N : {n m : } (n <= m) n ≤N m
<=Is≤N {n} {m} thm with orderIsTotal n m
<=Is≤N {n} {m} thm | inl (inl n<m) = inl n<m
<=Is≤N {n} {.n} thm | inr refl = inr refl
<=Is≤N {zero} {zero} thm | inl (inr m<n) = inl m<n
<=Is≤N {zero} {succ m} thm | inl (inr m<n) = inl (succIsPositive m)
<=Is≤N {succ n} {zero} () | inl (inr m<n)
<=Is≤N {succ n} {succ m} (os thm) | inl (inr m<n) with <=Is≤N thm
<=Is≤N {succ n} {succ m} (os thm) | inl (inr m<n) | inl x = inl (succPreservesInequality x)
<=Is≤N {succ n} {succ m} (os thm) | inl (inr m<n) | inr x rewrite x = inr refl
<=Is≤N {succ n} {succ m} (o' thm) | inl (inr m<n) with <=Is≤N thm
<=Is≤N {succ n} {succ m} (o' thm) | inl (inr m<n) | inl x with a<SuccA m
... | m<sm = inl (lessTransitive x m<sm)
<=Is≤N {succ n} {succ m} (o' thm) | inl (inr m<n) | inr x rewrite x = inl (a<SuccA m)
succIsNotEqual : {n : } (succ n n) False
succIsNotEqual {zero} pr = naughtE (equalityCommutative pr)
succIsNotEqual {succ n} pr = succIsNotEqual (succInjective pr)
succIsNotLess : {n : } (succ n <N n) False
succIsNotLess {zero} (le x proof) = naughtE (equalityCommutative proof)
succIsNotLess {succ n} pr = succIsNotLess (canRemoveSuccFrom<N pr)
noSN<=N : {n : } (succ n) <= n False
noSN<=N {n} thm with <=Is≤N thm
noSN<=N {n} thm | inl x = succIsNotLess x
noSN<=N {n} thm | inr x = succIsNotEqual x
no5<=4 : 5 <= 4 False
no5<=4 th = noSN<=N {4} th
_<?=_ : {a : _} {X : Set a} {n m : } (n <= m) Vec X m Vec X n
oz <?= xs = xs
os th <?= (x ,- xs) = x ,- (th <?= xs)
o' th <?= (x ,- xs) = th <?= xs
vMap<?=Fact : {a b : _} {X : Set a} {Y : Set b} (f : X Y) {n m : } (th : n <= m) (xs : Vec X m) vecMap f (th <?= xs) (th <?= vecMap f xs)
vMap<?=Fact f oz xs = refl
vMap<?=Fact f (os th) (x ,- xs) rewrite vMap<?=Fact f th xs = refl
vMap<?=Fact f (o' th) (x ,- xs) rewrite vMap<?=Fact f th xs = refl
oi : {n : } n <= n
oi {zero} = oz
oi {succ n} = os oi
oe : {n : } 0 <= n
oe {zero} = oz
oe {succ n} = o' oe
oeUnique : {n : } (th : 0 <= n) th oe
oeUnique oz = refl
oeUnique (o' i) rewrite oeUnique i = refl
oTooBig : {n m : } m ≤N n succ n <= m False
oTooBig {n} {m} m<=n th with <=Is≤N th
oTooBig {n} {m} (inl m<n) th | inl x = succIsNotLess (lessTransitive x m<n)
oTooBig {n} {m} (inr m=n) th | inl x rewrite m=n = succIsNotLess x
oTooBig {n} {m} (inl m<n) th | inr x rewrite (equalityCommutative x) = succIsNotLess m<n
oTooBig {n} {m} (inr m=n) th | inr x rewrite m=n = succIsNotEqual x
oiUnique : {n : } (th : n <= n) (th oi)
oiUnique oz = refl
oiUnique (os th) rewrite oiUnique th = refl
oiUnique {succ n} (o' th) with oTooBig {n} {n} (inr refl) th
... | bl = exFalso bl
id-<?= : {a : _} {X : Set a} {n : } (xs : Vec X n) (oi <?= xs) xs
id-<?= [] = refl
id-<?= (x ,- xs) rewrite id-<?= xs = refl
_o>>_ : {p n m : } p <= n n <= m p <= m
oz o>> th' = th'
os th o>> os th' = os (th o>> th')
os th o>> o' th' = os (o' th o>> th')
o' th o>> os th' = o' (th o>> th')
o' th o>> o' th' = o' (o' th o>> th')
eqHelper : {a : _} {X : Set a} (x : X) {n : } {r s : Vec X n} (pr : r s) (x ,- r) (x ,- s)
eqHelper x pr rewrite pr = refl
{-
cp-<?= : {p n m : } (th : p <= n) (th' : n <= m) → {a : _} {X : Set a} (xs : Vec X m) → ((th o>> th') <?= xs) ≡ (th <?= (th' <?= xs))
cp-<?= oz oz [] = refl
cp-<?= oz (o' th') xs = refl
cp-<?= (os th) (os th') (x ,- xs) = eqHelper x (cp-<?= th th' xs)
cp-<?= (o' th) (os th') (x ,- xs) = cp-<?= th th' xs
cp-<?= (o' th) (o' th') (x ,- xs) = cp-<?= (o' th) th' xs
cp-<?= (os th) (o' th') (x ,- xs) = {!!}
idThen-o>> : {n m : } (th : n <= m) → (oi o>> th) ≡ th
idThen-o>> oz = refl
idThen-o>> (os th) = applyEquality os (idThen-o>> th)
idThen-o>> {zero} (o' th) = refl
idThen-o>> {succ n} (o' th) with idThen-o>> th
... | bl = {!!}
idAfter-o>> : {n m : } (th : n <= m) → (th o>> oi) ≡ th
idAfter-o>> th = {!!}
-}

27
WellFoundedInduction.agda Normal file
View File

@@ -0,0 +1,27 @@
--{-# OPTIONS --safe --warning=error #-}
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Functions
module WellFoundedInduction where
data Accessible {a} {b} {A} (_<_ : Rel {a} {b} A) (x : A) : Set (lsuc a b) where
access : ( y y < x Accessible _<_ y) Accessible _<_ x
WellFounded : {a b : _} {A} Rel {a} {b} A Set (lsuc a b)
WellFounded {a} _<_ = x Accessible _<_ x
foldAcc : {a b c : _} {A : Set a} {_<_ : Rel {a} {c} A}
(P : A Set b)
( x ( y y < x P y) P x)
z Accessible _<_ z
P z
foldAcc {a} {b} {c} {A} {_<_} P inductionProof = go
where
go : (z : A) (Accessible _<_ z) P z
go z (access prf) = inductionProof z (λ y yLessZ go y (prf y yLessZ))
rec : {a b c : _} {A : Set a} {_<_ : Rel {a} {c} A}
WellFounded _<_
(P : A Set b)
( x ( y y < x P y) P x) ( z P z)
rec wf P inductionProof z = foldAcc P inductionProof _ (wf z)