Tidy up groups (#64)

This commit is contained in:
Patrick Stevens
2019-11-03 17:12:48 +00:00
committed by GitHub
parent e4daab7153
commit d95f510cdd
42 changed files with 1438 additions and 1038 deletions

View File

@@ -2,26 +2,19 @@
open import LogicalFormulae
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Numbers.Naturals
open import Numbers.Naturals.Naturals
open import Numbers.Naturals.Order
open import Vectors
open import Semirings.Definition
open import Categories.Definition
module Categories.Category where
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) }
dual record { objects = objects ; arrows = arrows ; id = id ; _∘_ = _∘_ ; rightId = rightId ; leftId = leftId ; compositionAssociative = 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 }
@@ -55,7 +48,7 @@ NatPreorder = record { objects = ; arrows = λ m n → m ≤N n ; id = λ x
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) }
NatMonoid = record { objects = True ; arrows = λ _ _ ; id = λ x 0 ; _∘_ = λ f g f +N g ; rightId = λ f refl ; leftId = λ f Semiring.sumZeroRight Semiring f ; associative = λ a b c Semiring.+Associative Semiring 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)) }

View File

@@ -0,0 +1,20 @@
{-# OPTIONS --warning=error --safe --without-K #-}
open import LogicalFormulae
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Numbers.Naturals.Naturals
open import Numbers.Naturals.Order
open import Vectors
open import Semirings.Definition
module Categories.Definition where
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
compositionAssociative : {x y z w : objects} (f : arrows z w) (g : arrows y z) (h : arrows x y) (f (g h)) (f g) h

30
Categories/Examples.agda Normal file
View File

@@ -0,0 +1,30 @@
{-# OPTIONS --warning=error --safe --without-K #-}
open import LogicalFormulae
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Numbers.Naturals.Naturals
open import Numbers.Naturals.Order
open import Vectors
open import Semirings.Definition
open import Categories.Definition
open import Groups.Definition
module Categories.Examples where
SET : {a : _} Category {lsuc a} {a}
Category.objects (SET {a}) = Set a
Category.arrows (SET {a}) = λ a b (a b)
Category.id (SET {a}) = λ x λ y y
Category._∘_ (SET {a}) = λ f g λ x f (g x)
Category.rightId (SET {a}) = λ f refl
Category.leftId (SET {a}) = λ f refl
Category.compositionAssociative (SET {a}) = λ f g h refl
GROUP : {a : _} Category {lsuc a} {a}
Category.objects (GROUP {a}) = Group {! !} {! !}
Category.arrows (GROUP {a}) = {! !}
Category.id (GROUP {a}) = {! !}
Category._∘_ (GROUP {a}) = {! !}
Category.rightId (GROUP {a}) = {! !}
Category.leftId (GROUP {a}) = {! !}
Category.compositionAssociative (GROUP {a}) = {! !}

View File

@@ -13,6 +13,11 @@ open import Numbers.Integers.Integers
open import Lists.Lists
open import Groups.Groups
open import Groups.Abelian.Lemmas
open import Groups.DirectSum.Definition
open import Groups.FiniteGroups.Definition
open import Groups.Homomorphisms.Lemmas
open import Groups.Isomorphisms.Lemmas
open import Groups.FinitePermutations
open import Groups.Lemmas
open import Groups.Groups2
@@ -62,4 +67,3 @@ open import Fields.CauchyCompletion.Group
open import Fields.CauchyCompletion.Ring
module Everything.Safe where

View File

@@ -4,14 +4,14 @@
open import Numbers.Primes.PrimeNumbers
open import Numbers.Primes.IntegerFactorisation
open import Numbers.Rationals
open import Numbers.RationalsLemmas
open import Numbers.Rationals.Definition
open import Numbers.Rationals.Lemmas
open import Logic.PropositionalLogic
open import Logic.PropositionalLogicExamples
open import Logic.PropositionalAxiomsTautology
open import Numbers.Modulo.IntegersModN
open import Numbers.Modulo.Group
open import Sets.FinSetWithK
@@ -23,5 +23,6 @@ open import Groups.Examples.ExampleSheet1
open import Groups.LectureNotes.Lecture1
open import LectureNotes.NumbersAndSets.Lecture1
open import LectureNotes.Groups.Lecture1
module Everything.WithK where

View File

@@ -7,6 +7,7 @@ open import Rings.Lemmas
open import Rings.Orders.Partial.Definition
open import Rings.Orders.Total.Definition
open import Groups.Definition
open import Groups.Lemmas
open import Groups.Groups
open import Fields.Fields
open import Sets.EquivalenceRelations

View File

@@ -9,6 +9,7 @@ open import Rings.Orders.Partial.Definition
open import Rings.Orders.Total.Definition
open import Groups.Definition
open import Groups.Groups
open import Groups.Lemmas
open import Fields.Fields
open import Sets.EquivalenceRelations
open import Sequences
@@ -123,7 +124,7 @@ abstract
... | f with totality 0G (am + inverse aN)
r | am-aN<e/4 | inl (inl 0<am-aN) = SetoidPartialOrder.<Transitive pOrder (<WellDefined (Equivalence.transitive eq (invContravariant additiveGroup) (Equivalence.transitive eq groupIsAbelian (+WellDefined (Equivalence.reflexive eq) (invTwice additiveGroup _)))) (Equivalence.reflexive eq) (lemm2' _ 0<am-aN)) 0<e/4
r | f | inl (inr x) = <WellDefined (Equivalence.transitive eq (invContravariant additiveGroup) (Equivalence.transitive eq groupIsAbelian (+WellDefined (Equivalence.reflexive eq) (invTwice additiveGroup _)))) (Equivalence.reflexive eq) f
r | am-aN<e/4 | inr 0=am-aN = <WellDefined (Equivalence.transitive eq (Equivalence.symmetric eq 0=am-aN) (Equivalence.transitive eq (Equivalence.transitive eq (Equivalence.transitive eq (Equivalence.transitive eq (Equivalence.symmetric eq (invIdentity additiveGroup)) (inverseWellDefined additiveGroup 0=am-aN)) (inverseWellDefined additiveGroup groupIsAbelian)) (invContravariant additiveGroup)) (+WellDefined (Equivalence.reflexive eq) (invTwice additiveGroup _)))) (Equivalence.reflexive eq) am-aN<e/4
r | am-aN<e/4 | inr 0=am-aN = <WellDefined (Equivalence.transitive eq (Equivalence.symmetric eq 0=am-aN) (Equivalence.transitive eq (Equivalence.transitive eq (Equivalence.transitive eq (Equivalence.transitive eq (Equivalence.symmetric eq (invIdent additiveGroup)) (inverseWellDefined additiveGroup 0=am-aN)) (inverseWellDefined additiveGroup groupIsAbelian)) (invContravariant additiveGroup)) (+WellDefined (Equivalence.reflexive eq) (invTwice additiveGroup _)))) (Equivalence.reflexive eq) am-aN<e/4
q : ((inverse (index (CauchyCompletion.elts a) m)) + (index (CauchyCompletion.elts a) (succ N) + e/2)) < (e/4 + e/2)
q = <WellDefined (Equivalence.symmetric eq +Associative) (Equivalence.reflexive eq) (orderRespectsAddition r e/2)

View File

@@ -8,6 +8,7 @@ open import Rings.Orders.Partial.Definition
open import Rings.Orders.Total.Definition
open import Groups.Definition
open import Groups.Groups
open import Groups.Lemmas
open import Fields.Fields
open import Fields.Orders.Total.Definition
open import Sets.EquivalenceRelations

View File

@@ -7,6 +7,8 @@ open import Rings.Lemmas
open import Rings.Orders.Partial.Definition
open import Rings.Orders.Total.Definition
open import Groups.Definition
open import Groups.Lemmas
open import Groups.Homomorphisms.Definition
open import Groups.Groups
open import Fields.Fields
open import Sets.EquivalenceRelations
@@ -50,7 +52,7 @@ CinvRight : {a : CauchyCompletion} → Setoid.__ cauchyCompletionSetoid (a +C
CinvRight {a} ε 0<e = 0 , ans
where
ans : {m : } (0 <N m) abs (index (apply _+_ (CauchyCompletion.elts (a +C (-C a))) (map inverse (CauchyCompletion.elts (injection 0G)))) m) < ε
ans {m} 0<m rewrite indexAndApply (CauchyCompletion.elts (a +C (-C a))) (map inverse (CauchyCompletion.elts (injection 0G))) _+_ {m} | indexAndApply (CauchyCompletion.elts a) (map inverse (CauchyCompletion.elts a)) _+_ {m} | equalityCommutative (mapAndIndex (CauchyCompletion.elts a) inverse m) | equalityCommutative (mapAndIndex (constSequence 0G) inverse m) | indexAndConst 0G m = <WellDefined (Equivalence.symmetric eq (Equivalence.transitive eq (absWellDefined _ _ (Equivalence.transitive eq (+WellDefined invRight (invIdentity (Ring.additiveGroup R))) identRight)) (identityOfIndiscernablesRight __ (Equivalence.reflexive eq) absZero))) (Equivalence.reflexive eq) 0<e
ans {m} 0<m rewrite indexAndApply (CauchyCompletion.elts (a +C (-C a))) (map inverse (CauchyCompletion.elts (injection 0G))) _+_ {m} | indexAndApply (CauchyCompletion.elts a) (map inverse (CauchyCompletion.elts a)) _+_ {m} | equalityCommutative (mapAndIndex (CauchyCompletion.elts a) inverse m) | equalityCommutative (mapAndIndex (constSequence 0G) inverse m) | indexAndConst 0G m = <WellDefined (Equivalence.symmetric eq (Equivalence.transitive eq (absWellDefined _ _ (Equivalence.transitive eq (+WellDefined invRight (invIdent (Ring.additiveGroup R))) identRight)) (identityOfIndiscernablesRight __ (Equivalence.reflexive eq) absZero))) (Equivalence.reflexive eq) 0<e
CGroup : Group cauchyCompletionSetoid _+C_
Group.+WellDefined CGroup {a} {b} {c} {d} x y = additionWellDefined {a} {c} {b} {d} x y

View File

@@ -8,6 +8,7 @@ open import Rings.Orders.Partial.Definition
open import Rings.Orders.Total.Definition
open import Groups.Definition
open import Groups.Groups
open import Groups.Lemmas
open import Fields.Fields
open import Sets.EquivalenceRelations
open import Sequences

View File

@@ -8,6 +8,7 @@ open import Rings.Orders.Partial.Definition
open import Rings.Orders.Total.Definition
open import Groups.Definition
open import Groups.Groups
open import Groups.Lemmas
open import Fields.Fields
open import Sets.EquivalenceRelations
open import Sequences

View File

@@ -8,6 +8,7 @@ open import Rings.Orders.Partial.Definition
open import Rings.Orders.Total.Definition
open import Groups.Definition
open import Groups.Groups
open import Groups.Lemmas
open import Fields.Fields
open import Sets.EquivalenceRelations
open import Sequences

View File

@@ -2,6 +2,7 @@
open import LogicalFormulae
open import Groups.Groups
open import Groups.Homomorphisms.Definition
open import Groups.Definition
open import Groups.Lemmas
open import Rings.Definition

View File

@@ -2,6 +2,7 @@
open import LogicalFormulae
open import Groups.Groups
open import Groups.Lemmas
open import Groups.Definition
open import Rings.Definition
open import Rings.Orders.Partial.Definition

View File

@@ -0,0 +1,17 @@
{-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae
open import Setoids.Setoids
open import Functions
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Numbers.Naturals.Naturals
open import Sets.FinSet
open import Groups.Definition
open import Sets.EquivalenceRelations
module Groups.Abelian.Definition where
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)

View File

@@ -0,0 +1,29 @@
{-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae
open import Setoids.Setoids
open import Functions
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Numbers.Naturals.Naturals
open import Sets.FinSet
open import Groups.Definition
open import Sets.EquivalenceRelations
open import Groups.Abelian.Definition
open import Groups.Homomorphisms.Definition
open import Groups.DirectSum.Definition
open import Groups.Subgroups.Definition
open import Groups.Isomorphisms.Definition
module Groups.Abelian.Lemmas where
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 (directSum underG underH))
AbelianGroup.commutative (directSumAbelianGroup {A = A} {B} G H) = AbelianGroup.commutative G ,, AbelianGroup.commutative H
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 Equivalence 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

View File

@@ -0,0 +1,22 @@
{-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae
open import Setoids.Setoids
open import Functions
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Numbers.Naturals.Naturals
open import Sets.FinSet
open import Groups.Definition
open import Sets.EquivalenceRelations
module Groups.DirectSum.Definition where
directSum : {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 (directSum {A = A} {B} g h) (s ,, t) (u ,, v) = Group.+WellDefined g s u ,, Group.+WellDefined h t v
Group.0G (directSum {A = A} {B} g h) = (Group.0G g ,, Group.0G h)
Group.inverse (directSum {A = A} {B} g h) (g1 ,, h1) = (Group.inverse g g1) ,, (Group.inverse h h1)
Group.+Associative (directSum {A = A} {B} g h) = Group.+Associative g ,, Group.+Associative h
Group.identRight (directSum {A = A} {B} g h) = Group.identRight g ,, Group.identRight h
Group.identLeft (directSum {A = A} {B} g h) = Group.identLeft g ,, Group.identLeft h
Group.invLeft (directSum {A = A} {B} g h) = Group.invLeft g ,, Group.invLeft h
Group.invRight (directSum {A = A} {B} g h) = Group.invRight g ,, Group.invRight h

View File

@@ -8,6 +8,12 @@ open import Numbers.Naturals.Naturals
open import Numbers.Integers.Integers
open import Sets.FinSet
open import Groups.Definition
open import Groups.Homomorphisms.Definition
open import Groups.Homomorphisms.Lemmas
open import Groups.Isomorphisms.Definition
open import Groups.Abelian.Definition
open import Groups.Subgroups.Definition
open import Groups.Lemmas
open import Groups.Groups
open import Rings.Definition
open import Rings.Lemmas

View File

@@ -0,0 +1,20 @@
{-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae
open import Setoids.Setoids
open import Functions
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Numbers.Naturals.Naturals
open import Sets.FinSet
open import Groups.Definition
open import Sets.EquivalenceRelations
module Groups.FiniteGroups.Definition where
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

View File

@@ -8,123 +8,14 @@ open import Numbers.Naturals.Naturals
open import Sets.FinSet
open import Groups.Definition
open import Sets.EquivalenceRelations
open import Groups.Homomorphisms.Definition
open import Groups.Lemmas
open import Groups.Homomorphisms.Lemmas
module Groups.Groups where
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)
open Setoid S
transitive = Equivalence.transitive (Setoid.eq S)
symmetric = Equivalence.symmetric (Setoid.eq S)
j : ((x ^-1) · x) · y (x ^-1) · (x · z)
j = Equivalence.transitive eq (symmetric (+Associative {x ^-1} {x} {y})) (+WellDefined ~refl pr)
k : ((x ^-1) · x) · y ((x ^-1) · x) · z
k = transitive j +Associative
l : 0G · y ((x ^-1) · x) · z
l = transitive (+WellDefined (symmetric invLeft) ~refl) k
m : 0G · y 0G · z
m = transitive l (+WellDefined invLeft ~refl)
n : y 0G · z
n = transitive (symmetric identLeft) m
o : y z
o = transitive n identLeft
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 identRight
where
open Group g renaming (inverse to _^-1)
open Setoid S
transitive = Equivalence.transitive (Setoid.eq S)
symmetric = Equivalence.symmetric (Setoid.eq S)
i : (y · x) · (x ^-1) (z · x) · (x ^-1)
i = +WellDefined pr ~refl
j : y · (x · (x ^-1)) (z · x) · (x ^-1)
j = transitive +Associative i
j' : y · 0G (z · x) · (x ^-1)
j' = transitive (+WellDefined ~refl (symmetric invRight)) j
k : y (z · x) · (x ^-1)
k = transitive (symmetric identRight) j'
l : y z · (x · (x ^-1))
l = transitive k (symmetric +Associative)
m : y z · 0G
m = transitive l (+WellDefined ~refl invRight)
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
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
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
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.0G 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)
open Setoid S
open Equivalence eq
i : y y · 0G
j : y · 0G y · (x · (x ^-1))
k : y · (x · (x ^-1)) (y · x) · (x ^-1)
l : (y · x) · (x ^-1) 0G · (x ^-1)
m : 0G · (x ^-1) x ^-1
i = symmetric identRight
j = +WellDefined ~refl (symmetric invRight)
k = +Associative
l = +WellDefined f ~refl
m = identLeft
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.0G 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)
open Setoid S
open Equivalence eq
i : y · (x · y) y · 0G
i' : y · (x · y) y
j : (y · x) · y y
k : (y · x) · y 0G · y
l : y · x 0G
i = +WellDefined ~refl f
i' = transitive i identRight
j = transitive (symmetric +Associative) i'
k = transitive j (symmetric identLeft)
l = groupsHaveRightCancellation G y (y · x) 0G 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)
open Setoid S
open Equivalence eq
fourWay+Associative : {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))
fourWay+Associative {S = S} {_·_} G {r} {s} {t} {u} = transitive p1 (transitive p2 p3)
where
@@ -152,153 +43,6 @@ fourWay+Associative'' {S = S} {_·_ = _·_} G = transitive +Associative (symmetr
open Setoid S
open Equivalence eq
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)
open Setoid S
open Equivalence eq
otherInv = (y ^-1) · (x ^-1)
many+Associatives : x · ((y · (y ^-1)) · (x ^-1)) (x · y) · ((y ^-1) · (x ^-1))
oneMult : (x · y) · otherInv x · (x ^-1)
many+Associatives = fourWay+Associative G
oneMult = symmetric (transitive (Group.+WellDefined G reflexive (transitive (symmetric (Group.identLeft G)) (Group.+WellDefined G (symmetric (Group.invRight G)) reflexive))) many+Associatives)
otherInvIsInverse : (x · y) · otherInv 0G
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.0G G)) (Group.0G G)
invIdentity {S = S} G = transitive (symmetric identLeft) invRight
where
open Group G
open Setoid S
open Equivalence eq
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.0G 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
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.0G 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
lemma : a · (inverse (inverse b)) 0G
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.0G 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 Equivalence 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.0G (directSumGroup {A = A} {B} g h) = (Group.0G g ,, Group.0G h)
Group.inverse (directSumGroup {A = A} {B} g h) (g1 ,, h1) = (Group.inverse g g1) ,, (Group.inverse h h1)
Group.+Associative (directSumGroup {A = A} {B} g h) = Group.+Associative g ,, Group.+Associative h
Group.identRight (directSumGroup {A = A} {B} g h) = Group.identRight g ,, Group.identRight h
Group.identLeft (directSumGroup {A = A} {B} g h) = Group.identLeft g ,, Group.identLeft 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.0G G)) (Group.0G H)
imageOfIdentityIsIdentity {S = S} {T = T} {_·A_ = _·A_} {_·B_ = _·B_} {G = G} {H = H} {f = f} hom = Equivalence.symmetric (Setoid.eq T) t
where
open Group H
open Setoid T
id2 : Setoid.__ S (Group.0G G) ((Group.0G G) ·A (Group.0G G))
id2 = Equivalence.symmetric (Setoid.eq S) (Group.identRight G)
r : f (Group.0G G) f (Group.0G G) ·B f (Group.0G G)
s : 0G ·B f (Group.0G G) f (Group.0G G) ·B f (Group.0G G)
t : 0G f (Group.0G G)
t = groupsHaveRightCancellation H (f (Group.0G G)) 0G (f (Group.0G G)) s
s = Equivalence.transitive (Setoid.eq T) identLeft r
r = Equivalence.transitive (Setoid.eq T) (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 = (Equivalence.transitive (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 Equivalence (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
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
@@ -311,7 +55,7 @@ quotientGroupSetoid {A = A} {S = S} {T = T} {_·A_ = _·A_} {_·B_ = _·B_} G {H
Equivalence.symmetric (Setoid.eq ansSetoid) {m} {n} pr = i
where
g : f (Group.inverse G (m ·A Group.inverse G n)) 0G
g = transitive (homRespectsInverse fHom {m ·A Group.inverse G n}) (transitive (inverseWellDefined H pr) (invIdentity H))
g = transitive (homRespectsInverse fHom {m ·A Group.inverse G n}) (transitive (inverseWellDefined H pr) (invIdent H))
h : f (Group.inverse G (Group.inverse G n) ·A Group.inverse G m) 0G
h = transitive (GroupHom.wellDefined fHom (Equivalence.symmetric (Setoid.eq S) (invContravariant G))) g
i : f (n ·A Group.inverse G m) 0G
@@ -382,20 +126,15 @@ Group.invLeft (quotientGroup {S = S} {T = T} {_·A_ = _·A_} G {H = H} {underf =
open Setoid T
open Equivalence eq
ans : f ((inverse x ·A x) ·A (inverse 0G)) (Group.0G H)
ans = transitive (GroupHom.wellDefined fHom (Equivalence.transitive (Setoid.eq S) (replaceGroupOp G (Equivalence.symmetric (Setoid.eq S) (Group.invLeft G)) (Equivalence.symmetric (Setoid.eq S) (invIdentity G)) (Equivalence.reflexive (Setoid.eq S)) ((Equivalence.reflexive (Setoid.eq S))) ((Equivalence.reflexive (Setoid.eq S)))) (identRight {0G}))) (imageOfIdentityIsIdentity fHom)
ans = transitive (GroupHom.wellDefined fHom (Equivalence.transitive (Setoid.eq S) (replaceGroupOp G (Equivalence.symmetric (Setoid.eq S) (Group.invLeft G)) (Equivalence.symmetric (Setoid.eq S) (invIdent G)) (Equivalence.reflexive (Setoid.eq S)) ((Equivalence.reflexive (Setoid.eq S))) ((Equivalence.reflexive (Setoid.eq S)))) (identRight {0G}))) (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
ans : f ((x ·A inverse x) ·A (inverse 0G)) (Group.0G H)
ans = transitive (GroupHom.wellDefined fHom (Equivalence.transitive (Setoid.eq S) (replaceGroupOp G (Equivalence.symmetric (Setoid.eq S) (Group.invRight G)) (Equivalence.symmetric (Setoid.eq S) (invIdentity G)) (Equivalence.reflexive (Setoid.eq S)) (Equivalence.reflexive (Setoid.eq S)) (Equivalence.reflexive (Setoid.eq S))) (identRight {0G}))) (imageOfIdentityIsIdentity fHom)
ans = transitive (GroupHom.wellDefined fHom (Equivalence.transitive (Setoid.eq S) (replaceGroupOp G (Equivalence.symmetric (Setoid.eq S) (Group.invRight G)) (Equivalence.symmetric (Setoid.eq S) (invIdent G)) (Equivalence.reflexive (Setoid.eq S)) (Equivalence.reflexive (Setoid.eq S)) (Equivalence.reflexive (Setoid.eq S))) (identRight {0G}))) (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 = {!!}
@@ -422,11 +161,3 @@ quotientIsSubgroup : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A
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 Equivalence 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

View File

@@ -10,6 +10,12 @@ open import Sets.FinSet
open import Functions
open import Sets.EquivalenceRelations
open import Numbers.Naturals.Naturals
open import Groups.Homomorphisms.Definition
open import Groups.Homomorphisms.Lemmas
open import Groups.Isomorphisms.Definition
open import Groups.Subgroups.Definition
open import Groups.Lemmas
open import Groups.Abelian.Definition
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
@@ -106,7 +112,7 @@ module Groups.Groups2 where
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.0G (groupKernelGroup G fHom) = kerOfElt (Group.0G 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)))
Group.inverse (groupKernelGroup {T = T} G {H = H} fHom) (kerOfElt x prX) = kerOfElt (Group.inverse G x) (transitive (homRespectsInverse fHom) (transitive (inverseWellDefined H prX) (invIdent H)))
where
open Setoid T
open Equivalence eq
@@ -161,4 +167,3 @@ module Groups.Groups2 where
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) = Equivalence.reflexive (Setoid.eq S)
GroupHom.wellDefined (identityHom G) = id

View File

@@ -0,0 +1,25 @@
{-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae
open import Setoids.Setoids
open import Functions
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Numbers.Naturals.Naturals
open import Sets.FinSet
open import Groups.Definition
open import Sets.EquivalenceRelations
module Groups.Homomorphisms.Definition where
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

View File

@@ -0,0 +1,42 @@
{-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae
open import Setoids.Setoids
open import Functions
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Numbers.Naturals.Naturals
open import Sets.FinSet
open import Groups.Definition
open import Sets.EquivalenceRelations
open import Groups.Homomorphisms.Definition
open import Groups.Lemmas
module Groups.Homomorphisms.Lemmas where
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.0G G)) (Group.0G H)
imageOfIdentityIsIdentity {S = S} {T = T} {_·A_ = _·A_} {_·B_ = _·B_} {G = G} {H = H} {f = f} hom = Equivalence.symmetric (Setoid.eq T) t
where
open Group H
open Setoid T
id2 : Setoid.__ S (Group.0G G) ((Group.0G G) ·A (Group.0G G))
id2 = Equivalence.symmetric (Setoid.eq S) (Group.identRight G)
r : f (Group.0G G) f (Group.0G G) ·B f (Group.0G G)
s : 0G ·B f (Group.0G G) f (Group.0G G) ·B f (Group.0G G)
t : 0G f (Group.0G G)
t = groupsHaveRightCancellation H (f (Group.0G G)) 0G (f (Group.0G G)) s
s = Equivalence.transitive (Setoid.eq T) identLeft r
r = Equivalence.transitive (Setoid.eq T) (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 = (Equivalence.transitive (Setoid.eq U)) (GroupHom.wellDefined gHom (GroupHom.groupHom fHom {x} {y}) ) (GroupHom.groupHom gHom {f x} {f y})
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

View File

@@ -0,0 +1,25 @@
{-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae
open import Setoids.Setoids
open import Functions
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Numbers.Naturals.Naturals
open import Sets.FinSet
open import Groups.Definition
open import Groups.Homomorphisms.Definition
open import Sets.EquivalenceRelations
module Groups.Isomorphisms.Definition where
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

View File

@@ -0,0 +1,38 @@
{-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae
open import Setoids.Setoids
open import Functions
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Numbers.Naturals.Naturals
open import Sets.FinSet
open import Groups.Definition
open import Sets.EquivalenceRelations
open import Groups.Isomorphisms.Definition
open import Groups.Homomorphisms.Definition
open import Groups.Homomorphisms.Lemmas
module Groups.Isomorphisms.Lemmas where
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 Equivalence (Setoid.eq U)
--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)))

View File

@@ -6,12 +6,15 @@ open import Functions
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Numbers.Naturals.Naturals
open import Numbers.Integers.Integers
open import Numbers.Rationals
open import Numbers.Rationals.Definition
open import Sets.FinSet
open import Groups.Definition
open import Groups.Groups
open import Groups.Abelian.Definition
open import Groups.FiniteGroups.Definition
open import Rings.Definition
open import Numbers.Modulo.IntegersModN
open import Numbers.Modulo.Group
open import Numbers.Modulo.Definition
open import Semirings.Definition
module Groups.LectureNotes.Lecture1 where

View File

@@ -5,35 +5,138 @@ open import Setoids.Setoids
open import Functions
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Numbers.Naturals.Naturals
open import Groups.Groups
open import Groups.Definition
open import Sets.EquivalenceRelations
module Groups.Lemmas where
invInv : {a b : _} {A : Set a} {_·_ : A A A} {S : Setoid {a} {b} A} (G : Group S _·_) {x : A} Setoid.__ S (Group.inverse G (Group.inverse G x)) x
invInv {S = S} G {x} = symmetric (transferToRight' G invRight)
where
open Setoid S
open Group G
open Equivalence eq
module Groups.Lemmas {a b : _} {A : Set a} {_·_ : A A A} {S : Setoid {a} {b} A} (G : Group S _·_) where
invIdent : {a b : _} {A : Set a} {_·_ : A A A} {S : Setoid {a} {b} A} (G : Group S _·_) Setoid.__ S (Group.inverse G (Group.0G G)) (Group.0G G)
invIdent {S = S} G = symmetric (transferToRight' G (Group.identLeft G))
where
open Setoid S
open Group G
open Equivalence eq
open Setoid S
open Group G
open Equivalence eq
swapInv : {a b : _} {A : Set a} {_+_ : A A A} {S : Setoid {a} {b} A} (G : Group S _+_) {x y : A} Setoid.__ S (Group.inverse G x) y Setoid.__ S x (Group.inverse G y)
swapInv {S = S} G {x} {y} -x=y = transitive (symmetric (invInv G)) (inverseWellDefined G -x=y)
where
open Setoid S
open Group G
open Equivalence eq
groupsHaveLeftCancellation : (x y z : A) (x · y) (x · z) y z
groupsHaveLeftCancellation x y z pr = o
where
_^-1 = inverse
j : ((x ^-1) · x) · y (x ^-1) · (x · z)
j = transitive (symmetric (+Associative {x ^-1} {x} {y})) (+WellDefined ~refl pr)
k : ((x ^-1) · x) · y ((x ^-1) · x) · z
k = transitive j +Associative
l : 0G · y ((x ^-1) · x) · z
l = transitive (+WellDefined (symmetric invLeft) ~refl) k
m : 0G · y 0G · z
m = transitive l (+WellDefined invLeft ~refl)
n : y 0G · z
n = transitive (symmetric identLeft) m
o : y z
o = transitive n identLeft
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.0G G))
identityIsUnique {S = S} {_·_} g thing fb = transitive (symmetric identLeft) (fb 0G)
where
open Group g renaming (inverse to _^-1)
open Setoid S
open Equivalence eq
groupsHaveRightCancellation : (x y z : A) (y · x) (z · x) y z
groupsHaveRightCancellation x y z pr = transitive m identRight
where
_^-1 = inverse
i : (y · x) · (x ^-1) (z · x) · (x ^-1)
i = +WellDefined pr ~refl
j : y · (x · (x ^-1)) (z · x) · (x ^-1)
j = transitive +Associative i
j' : y · 0G (z · x) · (x ^-1)
j' = transitive (+WellDefined ~refl (symmetric invRight)) j
k : y (z · x) · (x ^-1)
k = transitive (symmetric identRight) j'
l : y z · (x · (x ^-1))
l = transitive k (symmetric +Associative)
m : y z · 0G
m = transitive l (+WellDefined ~refl invRight)
rightInversesAreUnique : (x : A) (y : A) (y · x) 0G y (inverse x)
rightInversesAreUnique x y f = transitive i (transitive j (transitive k (transitive l m)))
where
_^-1 = inverse
i : y y · 0G
j : y · 0G y · (x · (x ^-1))
k : y · (x · (x ^-1)) (y · x) · (x ^-1)
l : (y · x) · (x ^-1) 0G · (x ^-1)
m : 0G · (x ^-1) x ^-1
i = symmetric identRight
j = +WellDefined ~refl (symmetric invRight)
k = +Associative
l = +WellDefined f ~refl
m = identLeft
leftInversesAreUnique : {x : A} {y : A} (x · y) 0G y (inverse x)
leftInversesAreUnique {x} {y} f = rightInversesAreUnique x y l
where
_^-1 = inverse
i : y · (x · y) y · 0G
i' : y · (x · y) y
j : (y · x) · y y
k : (y · x) · y 0G · y
l : y · x 0G
i = +WellDefined ~refl f
i' = transitive i identRight
j = transitive (symmetric +Associative) i'
k = transitive j (symmetric identLeft)
l = groupsHaveRightCancellation y (y · x) 0G k
invTwice : (x : A) (inverse (inverse x)) x
invTwice x = symmetric (rightInversesAreUnique (x ^-1) x invRight)
where
_^-1 = inverse
replaceGroupOp : {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 a~c b~d w~y x~z pr = transitive (symmetric (+WellDefined a~c b~d)) (transitive pr (+WellDefined w~y x~z))
replaceGroupOpRight : {a b c x : A} (Setoid.__ S a (b · c)) (Setoid.__ S c x) (Setoid.__ S a (b · x))
replaceGroupOpRight a~bc c~x = transitive a~bc (+WellDefined reflexive c~x)
inverseWellDefined : {x y : A} (x y) (inverse x) (inverse y)
inverseWellDefined {x} {y} x~y = groupsHaveRightCancellation x (inverse x) (inverse y) q
where
p : inverse x · x inverse y · y
p = transitive invLeft (symmetric invLeft)
q : inverse x · x inverse y · x
q = replaceGroupOpRight {_·_ (inverse x) x} {inverse y} {y} {x} p (symmetric x~y)
transferToRight : {a b : A} (a · (inverse b)) 0G a b
transferToRight {a} {b} ab-1 = transitive (symmetric (invTwice a)) (transitive u (invTwice b))
where
t : inverse a inverse b
t = symmetric (leftInversesAreUnique ab-1)
u : inverse (inverse a) inverse (inverse b)
u = inverseWellDefined t
transferToRight' : {a b : A} (a · b) 0G a (inverse b)
transferToRight' {a} {b} ab-1 = transferToRight lemma
where
lemma : a · (inverse (inverse b)) 0G
lemma = transitive (+WellDefined reflexive (invTwice b)) ab-1
transferToRight'' : {a b : A} Setoid.__ S a b (a · (inverse b)) 0G
transferToRight'' {a} {b} a~b = transitive (+WellDefined a~b reflexive) invRight
invInv : {x : A} (inverse (inverse x)) x
invInv {x} = symmetric (transferToRight' invRight)
invIdent : (inverse 0G) 0G
invIdent = symmetric (transferToRight' identLeft)
swapInv : {x y : A} (inverse x) y x (inverse y)
swapInv {x} {y} -x=y = transitive (symmetric invInv) (inverseWellDefined -x=y)
identityIsUnique : (e : A) ((b : A) ((b · e) b)) (e 0G)
identityIsUnique thing fb = transitive (symmetric identLeft) (fb 0G)
invContravariant : {x y : A} (Setoid.__ S (Group.inverse G (x · y)) ((Group.inverse G y) · (Group.inverse G x)))
invContravariant {x} {y} = ans
where
_^-1 = inverse
otherInv = (y ^-1) · (x ^-1)
many+Associatives : x · ((y · (y ^-1)) · (x ^-1)) (x · y) · ((y ^-1) · (x ^-1))
oneMult : (x · y) · otherInv x · (x ^-1)
many+Associatives = transitive +Associative (transitive (+WellDefined +Associative reflexive) (symmetric +Associative))
oneMult = symmetric (transitive (+WellDefined reflexive (transitive (symmetric identLeft) (+WellDefined (symmetric invRight) reflexive))) many+Associatives)
otherInvIsInverse : (x · y) · otherInv 0G
otherInvIsInverse = transitive oneMult invRight
ans : (x · y) ^-1 (y ^-1) · (x ^-1)
ans = symmetric (leftInversesAreUnique otherInvIsInverse)

View File

@@ -0,0 +1,18 @@
{-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae
open import Setoids.Setoids
open import Functions
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Numbers.Naturals.Naturals
open import Sets.FinSet
open import Groups.Definition
open import Groups.Homomorphisms.Definition
module Groups.Subgroups.Definition where
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

View File

@@ -0,0 +1,205 @@
{-# OPTIONS --warning=error --safe #-}
open import Functions
open import Sets.FinSet
open import LogicalFormulae
open import Groups.Definition
open import Groups.Groups
open import Groups.FiniteGroups.Definition
open import Groups.Abelian.Definition
open import Setoids.Setoids
open import Numbers.Naturals.Naturals
open import Numbers.Integers.Integers
open import Numbers.Rationals.Definition
open import Rings.Definition
open import Fields.FieldOfFractions.Setoid
open import Semirings.Definition
open import Numbers.Modulo.Definition
open import Numbers.Modulo.Group
module LectureNotes.Groups.Lecture1 where
-- Examples
groupExample1 : Group (reflSetoid ) (_+Z_)
groupExample1 = Group
groupExample2 : Group (fieldOfFractionsSetoid IntDom) (_+Q_)
groupExample2 = Ring.additiveGroup Ring
groupExample3 : Group (reflSetoid ) (_-Z_) False
groupExample3 record { +Associative = multAssoc } with multAssoc {nonneg 3} {nonneg 2} {nonneg 1}
groupExample3 record { +WellDefined = wellDefined } | ()
negSuccInjective : {a b : } (negSucc a negSucc b) a b
negSuccInjective {a} {.a} refl = refl
nonnegInjective : {a b : } (nonneg a nonneg b) a b
nonnegInjective {a} {.a} refl = refl
integersTimesNotGroup : Group (reflSetoid ) (_*Z_) False
integersTimesNotGroup record { +WellDefined = wellDefined ; 0G = (nonneg zero) ; inverse = inverse ; +Associative = multAssoc ; identRight = multIdentRight ; identLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } with multIdentLeft {negSucc 1}
... | ()
integersTimesNotGroup record { +WellDefined = wellDefined ; 0G = (nonneg (succ zero)) ; inverse = inverse ; +Associative = multAssoc ; identRight = multIdentRight ; identLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } with invLeft {nonneg zero}
... | bl with inverse (nonneg zero)
integersTimesNotGroup record { +WellDefined = wellDefined ; 0G = (nonneg (succ zero)) ; inverse = inverse ; +Associative = multAssoc ; identRight = multIdentRight ; identLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } | () | nonneg zero
integersTimesNotGroup record { +WellDefined = wellDefined ; 0G = (nonneg (succ zero)) ; inverse = inverse ; +Associative = multAssoc ; identRight = multIdentRight ; identLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } | p | nonneg (succ x) = naughtE (nonnegInjective (transitivity (applyEquality nonneg (equalityCommutative (Semiring.productZeroRight Semiring x))) p))
integersTimesNotGroup record { +WellDefined = wellDefined ; 0G = (nonneg (succ zero)) ; inverse = inverse ; +Associative = multAssoc ; identRight = multIdentRight ; identLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } | () | negSucc x
integersTimesNotGroup record { +WellDefined = wellDefined ; 0G = (nonneg (succ (succ x))) ; inverse = inverse ; +Associative = multAssoc ; identRight = multIdentRight ; identLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } with succInjective (negSuccInjective (multIdentLeft {negSucc 1}))
... | ()
integersTimesNotGroup record { +WellDefined = wellDefined ; 0G = (negSucc x) ; inverse = inverse ; +Associative = multAssoc ; identRight = multIdentRight ; identLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } with multIdentLeft {nonneg 2}
integersTimesNotGroup record { +WellDefined = wellDefined ; 0G = (negSucc x) ; inverse = inverse ; +Associative = multAssoc ; identRight = multIdentRight ; identLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } | ()
-- TODO: Q is not a group with *Q
-- TODO: Q without 0 is a group with *Q
-- TODO: {1, -1} is a group with *
nIsGroup : (n : ) (0<n : 0 <N n) _
nIsGroup n pr = nGroup n pr
-- 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.0G weirdGroup = e
Group.inverse weirdGroup t = t
Group.+Associative weirdGroup {r} {s} {t} = +WAssoc {r} {s} {t}
Group.identRight weirdGroup {e} = refl
Group.identRight weirdGroup {a} = refl
Group.identRight weirdGroup {b} = refl
Group.identRight weirdGroup {c} = refl
Group.identLeft weirdGroup {e} = refl
Group.identLeft weirdGroup {a} = refl
Group.identLeft weirdGroup {b} = refl
Group.identLeft 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

View File

@@ -5,6 +5,7 @@ open import Numbers.Naturals.Naturals
open import Numbers.Integers.Definition
open import Semirings.Definition
open import Groups.Groups
open import Groups.Abelian.Definition
open import Groups.Definition
open import Setoids.Setoids

View File

@@ -0,0 +1,102 @@
{-# OPTIONS --safe --warning=error #-}
-- These are explicitly with-K, because we currently encode an element of Zn as
-- a natural together with a proof that it is small.
open import LogicalFormulae
open import Groups.Definition
open import Groups.Groups
open import Groups.Abelian.Definition
open import Groups.FiniteGroups.Definition
open import Numbers.Naturals.Naturals
open import Numbers.Naturals.Addition -- TODO remove this dependency
open import Numbers.Primes.PrimeNumbers
open import Setoids.Setoids
open import Sets.FinSet
open import Sets.FinSetWithK
open import Functions
open import Numbers.Naturals.WithK
open import Semirings.Definition
open import Orders
open import Numbers.Modulo.Definition
module Numbers.Modulo.Addition where
cancelSumFromInequality : {a b c : } a +N b <N a +N c b <N c
cancelSumFromInequality {a} {b} {c} (le x proof) = le x help
where
help : succ x +N b c
help rewrite Semiring.+Associative Semiring (succ x) a b | Semiring.commutative Semiring x a | equalityCommutative (Semiring.+Associative Semiring (succ a) x b) | Semiring.commutative Semiring a (x +N b) | Semiring.commutative Semiring (succ (x +N b)) a = canSubtractFromEqualityLeft {a} proof
_+n_ : {n : } {pr : 0 <N n} n n pr n n pr n n pr
_+n_ {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 _ _ (Semiring.commutative Semiring a 0)
plusZnIdentityRight {succ x} {_} record { x = a ; xLess = aLess } | inl (inr sx<a) = exFalso (f aLess sx<a)
where
f : (aL : a <N succ x) (sx<a : succ x <N a +N 0) False
f aL sx<a rewrite Semiring.commutative Semiring a 0 = 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 Semiring.commutative Semiring a 0 | a=sx = TotalOrder.irreflexive TotalOrder 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 (TotalOrder.irreflexive TotalOrder (TotalOrder.<Transitive TotalOrder succn<x xLess))
plusZnIdentityLeft {succ n} {pr} record { x = x ; xLess = xLess } | inr x=succn rewrite x=succn = exFalso (TotalOrder.irreflexive TotalOrder xLess)
subLemma : {a b c : } (pr1 : a <N b) (pr2 : a <N c) (pr : b c) (subtractionNResult.result (-N (inl pr1))) (subtractionNResult.result (-N (inl pr2)))
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 _ _ (Semiring.commutative Semiring a b)
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inl (inl a+b<sn) | inl (inr sn<b+a) = exFalso (f a+b<sn sn<b+a)
where
f : (a +N b) <N succ n succ n <N b +N a False
f pr1 pr2 rewrite Semiring.commutative Semiring b a = TotalOrder.irreflexive TotalOrder (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 Semiring.commutative Semiring b a | eq = TotalOrder.irreflexive TotalOrder pr1
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inl (inr sn<a+b) with 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 Semiring.commutative Semiring b a = TotalOrder.irreflexive TotalOrder (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 (Semiring.commutative Semiring a b))
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inl (inr sn<a+b) | inr sn=b+a = exFalso (f sn<a+b sn=b+a)
where
f : succ n <N a +N b b +N a succ n False
f pr eq rewrite Semiring.commutative Semiring b a | eq = TotalOrder.irreflexive TotalOrder pr
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inr sn=a+b with 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 Semiring.commutative Semiring b a | sn=a+b = TotalOrder.irreflexive TotalOrder b+a<sn
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inr sn=a+b | inl (inr sn<b+a) = exFalso f
where
f : False
f rewrite Semiring.commutative Semiring b a | sn=a+b = TotalOrder.irreflexive TotalOrder sn<b+a
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inr sn=a+b | inr sn=b+a = equalityZn _ _ refl

View File

@@ -0,0 +1,28 @@
{-# OPTIONS --safe --warning=error #-}
-- These are explicitly with-K, because we currently encode an element of Zn as
-- a natural together with a proof that it is small.
open import LogicalFormulae
open import Groups.Definition
open import Groups.Groups
open import Groups.Abelian.Definition
open import Groups.FiniteGroups.Definition
open import Numbers.Naturals.Naturals
open import Numbers.Naturals.Addition -- TODO remove this dependency
open import Numbers.Primes.PrimeNumbers
open import Setoids.Setoids
open import Sets.FinSet
open import Sets.FinSetWithK
open import Functions
open import Numbers.Naturals.WithK
open import Semirings.Definition
open import Orders
module Numbers.Modulo.Definition where
record n (n : ) (pr : 0 <N n) : Set where
field
x :
xLess : x <N n
equalityZn : {n : } {pr : 0 <N n} (a b : n n pr) (n.x a n.x b) a b
equalityZn record { x = a ; xLess = aLess } record { x = .a ; xLess = bLess } refl rewrite <NRefl aLess bLess = refl

622
Numbers/Modulo/Group.agda Normal file
View File

@@ -0,0 +1,622 @@
{-# OPTIONS --safe --warning=error #-}
-- These are explicitly with-K, because we currently encode an element of Zn as
-- a natural together with a proof that it is small.
open import LogicalFormulae
open import Groups.Definition
open import Groups.Groups
open import Groups.Abelian.Definition
open import Groups.FiniteGroups.Definition
open import Numbers.Naturals.Naturals
open import Numbers.Naturals.Addition -- TODO remove this dependency
open import Numbers.Primes.PrimeNumbers
open import Setoids.Setoids
open import Sets.FinSet
open import Sets.FinSetWithK
open import Functions
open import Numbers.Naturals.WithK
open import Semirings.Definition
open import Orders
open import Numbers.Modulo.Definition
open import Numbers.Modulo.Addition
module Numbers.Modulo.Group where
help30 : {a b c n : } (c <N n) (a +N b n) (n<b+c : n <N b +N c) (n <N a +N subtractionNResult.result (-N (inl n<b+c))) False
help30 {a} {b} {c} {n} c<n a+b=n n<b+c x = TotalOrder.irreflexive TotalOrder (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) (equalityCommutative (Semiring.+Associative Semiring a _ n))
pr2 : n +N n <N a +N (b +N c)
pr2 = identityOfIndiscernablesRight _<N_ pr (applyEquality (a +N_) (addMinus (inl n<b+c)))
pr3 : n +N n <N (a +N b) +N c
pr3 rewrite equalityCommutative (Semiring.+Associative Semiring a b c) = pr2
pr4 : n +N n <N c +N n
pr4 rewrite Semiring.commutative Semiring c n = identityOfIndiscernablesRight _<N_ pr3 (applyEquality (_+N c) a+b=n)
pr5 : n <N c
pr5 = subtractionPreservesInequality n pr4
help31 : {a b c n : } (a +N b n) (n<b+c : n <N b +N c) (a +N subtractionNResult.result (-N (inl n<b+c))) c
help31 {a} {b} {c} {n} a+b=n n<b+c = canSubtractFromEqualityLeft pr2
where
pr1 : a +N (n +N subtractionNResult.result (-N (inl n<b+c))) n +N c
pr1 rewrite addMinus' (inl n<b+c) | Semiring.+Associative Semiring a b c | a+b=n = refl
pr2 : n +N (a +N subtractionNResult.result (-N (inl n<b+c))) n +N c
pr2 = identityOfIndiscernablesLeft _≡_ pr1 (lemm a n (subtractionNResult.result (-N (inl n<b+c))))
where
lemm : (a b c : ) a +N (b +N c) b +N (a +N c)
lemm a b c rewrite Semiring.+Associative Semiring a b c | Semiring.commutative Semiring a b | equalityCommutative (Semiring.+Associative Semiring b a c) = refl
help7 : {a b c n : } b +N c n a <N n (n<a+b : n <N a +N b) (subtractionNResult.result (-N (inl n<a+b)) +N c n) False
help7 {a} {b} {c} {n} b+c=n a<n n<a+b x = TotalOrder.irreflexive TotalOrder (identityOfIndiscernablesLeft _<N_ a<n pr5)
where
pr : (n +N subtractionNResult.result (-N (inl n<a+b))) +N c n +N n
pr = identityOfIndiscernablesLeft _≡_ (applyEquality (n +N_) x) (Semiring.+Associative Semiring n _ c)
pr2 : (a +N b) +N c n +N n
pr2 = identityOfIndiscernablesLeft _≡_ pr (applyEquality (_+N c) (addMinus' (inl n<a+b)))
pr3 : a +N (b +N c) n +N n
pr3 rewrite Semiring.+Associative Semiring a b c = pr2
pr4 : a +N n n +N n
pr4 = identityOfIndiscernablesLeft _≡_ pr3 (applyEquality (a +N_) b+c=n)
pr5 : a n
pr5 = canSubtractFromEqualityRight pr4
help29 : {a b c n : } (c <N n) (n<b+c : n <N b +N c) (a +N subtractionNResult.result (-N (inl n<b+c))) n a +N b n False
help29 {a} {b} {c} {n} c<n n<b+c pr a+b=n = TotalOrder.irreflexive TotalOrder (identityOfIndiscernablesLeft _<N_ c<n p4)
where
p1 : a +N (subtractionNResult.result (-N (inl n<b+c)) +N n) n +N n
p1 = identityOfIndiscernablesLeft _≡_ (applyEquality (_+N n) pr) (equalityCommutative (Semiring.+Associative Semiring a _ n))
p2 : (a +N b) +N c n +N n
p2 rewrite equalityCommutative (Semiring.+Associative Semiring a b c) = identityOfIndiscernablesLeft _≡_ p1 (applyEquality (a +N_) (addMinus (inl n<b+c)))
p3 : n +N c n +N n
p3 = identityOfIndiscernablesLeft _≡_ p2 (applyEquality (_+N c) a+b=n)
p4 : c n
p4 = canSubtractFromEqualityLeft p3
help28 : {a b c n : } (n<a+'b+c : n <N a +N (b +N c)) (a +N b n) subtractionNResult.result (-N (inl n<a+'b+c)) c
help28 {a} {b} {c} {n} pr a+b=n = canSubtractFromEqualityLeft p2
where
p1 : a +N (b +N c) n +N c
p1 rewrite Semiring.+Associative Semiring a b c | a+b=n = refl
p2 : n +N subtractionNResult.result (-N (inl pr)) n +N c
p2 = identityOfIndiscernablesLeft _≡_ p1 (equalityCommutative (addMinus' (inl pr)))
help27 : {a b c n : } (a +N b succ n) (a +N (b +N c) <N succ n) a +N (b +N c) c
help27 {a} {b} {zero} {n} a+b=sn a+b+c<sn rewrite Semiring.commutative Semiring b 0 | a+b=sn = exFalso (TotalOrder.irreflexive TotalOrder a+b+c<sn)
help27 {a} {b} {succ c} {n} a+b=sn a+b+c<sn rewrite Semiring.+Associative Semiring a b (succ c) | a+b=sn = exFalso (cannotAddAndEnlarge' a+b+c<sn)
help26 : {a b c n : } (a +N b n) (a +N (b +N c) n) 0 c
help26 {a} {b} {c} {n} a+b=n a+b+c=n rewrite Semiring.+Associative Semiring a b c | a+b=n | Semiring.commutative Semiring n c = equalityCommutative (canSubtractFromEqualityRight a+b+c=n)
help25 : {a b c n : } (a +N b n) (b +N c n) (a +N 0 c)
help25 {a} {b} {c} {n} a+b=n b+c=n rewrite Semiring.commutative Semiring a 0 | Semiring.commutative Semiring a b | equalityCommutative a+b=n = equalityCommutative (canSubtractFromEqualityLeft b+c=n)
help24 : {a n : } (a <N n) (n <N a +N 0) False
help24 {a} {n} a<n n<a+0 rewrite Semiring.commutative Semiring a 0 = TotalOrder.irreflexive TotalOrder (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 Semiring.commutative Semiring a 0 | a+0=n = TotalOrder.irreflexive TotalOrder a<n
help22 : {a b c n : } (a +N b n) (c n) (n<b+c : n <N b +N c) (n <N a +N subtractionNResult.result (-N (inl n<b+c))) False
help22 {a} {b} {c} {.c} a+b=n refl n<b+c pr = TotalOrder.irreflexive TotalOrder (identityOfIndiscernablesRight _<N_ pr4 a+b=n)
where
pr1 : c +N c <N a +N (subtractionNResult.result (-N (inl n<b+c)) +N c)
pr1 = identityOfIndiscernablesRight _<N_ (additionPreservesInequality c pr) (equalityCommutative (Semiring.+Associative Semiring a _ c))
pr2 : c +N c <N a +N (b +N c)
pr2 = identityOfIndiscernablesRight _<N_ pr1 (applyEquality (a +N_) (addMinus (inl n<b+c)))
pr3 : c +N c <N (a +N b) +N c
pr3 = identityOfIndiscernablesRight _<N_ pr2 (Semiring.+Associative Semiring a b c)
pr4 : c <N a +N b
pr4 = subtractionPreservesInequality c pr3
help21 : {a b c n : } (a +N b n) (b +N c n) (c n) (a <N n) False
help21 {a} {b} {c} {.c} a+b=n b+c=n refl a<n = TotalOrder.irreflexive TotalOrder (identityOfIndiscernablesLeft _<N_ a<n a=c)
where
b=0 : b 0
a=c : a c
a=c = identityOfIndiscernablesLeft _≡_ a+b=n lemm
where
lemm : a +N b a
lemm rewrite b=0 | Semiring.commutative Semiring a 0 = refl
b=0' : (b c : ) (b +N c c) b 0
b=0' zero c b+c=c = refl
b=0' (succ b) zero b+c=c = exFalso (naughtE (equalityCommutative b+c=c))
b=0' (succ b) (succ c) b+c=c = b=0' (succ b) c bl2
where
bl2 : succ b +N c c
bl2 rewrite Semiring.commutative Semiring b c | Semiring.commutative Semiring (succ c) b = succInjective b+c=c
b=0 = b=0' b c b+c=n
help20 : {a b c n : } (c n) (a +N b n) (n<b+c : n <N b +N c) (a +N subtractionNResult.result (-N (inl n<b+c)) <N n) False
help20 {a} {b} {c} {n} c=n a+b=n n<b+c pr = TotalOrder.irreflexive TotalOrder (identityOfIndiscernablesLeft _<N_ pr5 c=n)
where
pr1 : a +N (subtractionNResult.result (-N (inl n<b+c)) +N n) <N n +N n
pr1 = identityOfIndiscernablesLeft _<N_ (additionPreservesInequality n pr) (equalityCommutative (Semiring.+Associative Semiring a _ n))
pr2 : a +N (b +N c) <N n +N n
pr2 = identityOfIndiscernablesLeft _<N_ pr1 (applyEquality (a +N_) (addMinus (inl n<b+c)))
pr3 : (a +N b) +N c <N n +N n
pr3 rewrite equalityCommutative (Semiring.+Associative Semiring a b c) = pr2
pr4 : c +N n <N n +N n
pr4 = identityOfIndiscernablesLeft _<N_ pr3 (transitivity (applyEquality (_+N c) a+b=n) (Semiring.commutative Semiring n c))
pr5 : c <N n
pr5 = subtractionPreservesInequality n pr4
help19 : {a b c n : } (b+c<n : b +N c <N n) (n<a+b : n <N a +N b) (a <N n) (subtractionNResult.result (-N (inl n<a+b)) +N c n) False
help19 {a} {b} {c} {n} b+c<n n<a+b a<n pr = TotalOrder.irreflexive TotalOrder (identityOfIndiscernablesLeft _<N_ r q')
where
p : (n +N subtractionNResult.result (-N (inl n<a+b))) +N c n +N n
p = identityOfIndiscernablesLeft _≡_ (applyEquality (n +N_ ) pr) (Semiring.+Associative Semiring n _ c)
q : (a +N b) +N c n +N n
q = identityOfIndiscernablesLeft _≡_ p (applyEquality (_+N c) (addMinus' (inl n<a+b)))
q' : a +N (b +N c) n +N n
q' = identityOfIndiscernablesLeft _≡_ q (equalityCommutative (Semiring.+Associative Semiring a b c))
r : a +N (b +N c) <N n +N n
r = addStrongInequalities a<n b+c<n
help18 : {a b c n : } (b+c<n : b +N c <N n) (n<a+b : n <N a +N b) (a <N n) (n <N subtractionNResult.result (-N (inl n<a+b)) +N c) False
help18 {a} {b} {c} {n} b+c<n n<a+b a<n pr = TotalOrder.irreflexive TotalOrder (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) (Semiring.+Associative Semiring n _ c)
p' : n +N n <N (a +N b) +N c
p' = identityOfIndiscernablesRight _<N_ p (applyEquality (_+N c) (addMinus' (inl n<a+b)))
p2 : n +N n <N a +N (b +N c)
p2 = identityOfIndiscernablesRight _<N_ p' (equalityCommutative (Semiring.+Associative Semiring a b c))
p3 : n +N n <N a +N n
p3 = 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 = TotalOrder.irreflexive TotalOrder (identityOfIndiscernablesLeft _<N_ pr1'' pr3)
where
pr1' : a +N (subtractionNResult.result (-N (inl n<b+c)) +N n) <N n +N n
pr1' = identityOfIndiscernablesLeft _<N_ (additionPreservesInequality n p1) (equalityCommutative (Semiring.+Associative Semiring a _ n))
pr1'' : a +N (b +N c) <N n +N n
pr1'' = identityOfIndiscernablesLeft _<N_ pr1' (applyEquality (a +N_) (addMinus (inl n<b+c)))
pr2' : (n +N subtractionNResult.result (-N (inl n<a+b))) +N c n +N n
pr2' = identityOfIndiscernablesLeft _≡_ (applyEquality (n +N_) p2) (Semiring.+Associative Semiring n _ c)
pr2'' : (a +N b) +N c n +N n
pr2'' = identityOfIndiscernablesLeft _≡_ pr2' (applyEquality (_+N c) (addMinus' (inl n<a+b)))
pr3 : a +N (b +N c) n +N n
pr3 = identityOfIndiscernablesLeft _≡_ pr2'' (equalityCommutative (Semiring.+Associative Semiring a b c))
help16 : {a b c n : } (n<b+c : n <N b +N c) (n<a+b : n <N a +N b) (a +N subtractionNResult.result (-N (inl n<b+c))) <N n (pr : n <N subtractionNResult.result (-N (inl n<a+b)) +N c) a +N subtractionNResult.result (-N (inl n<b+c)) subtractionNResult.result (-N (inl pr))
help16 {a} {b} {c} {n} n<b+c n<a+b pr1 pr2 = exFalso (TotalOrder.irreflexive TotalOrder (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) (equalityCommutative (Semiring.+Associative Semiring a _ n))
pr1'' : a +N (b +N c) <N n +N n
pr1'' = identityOfIndiscernablesLeft _<N_ pr1' (applyEquality (a +N_) (addMinus (inl n<b+c)))
pr2' : n +N n <N (n +N subtractionNResult.result (-N (inl n<a+b))) +N c
pr2' = identityOfIndiscernablesRight _<N_ (additionPreservesInequalityOnLeft n pr2) (Semiring.+Associative Semiring n _ c)
pr2'' : n +N n <N (a +N b) +N c
pr2'' = identityOfIndiscernablesRight _<N_ pr2' (applyEquality (_+N c) (addMinus' (inl n<a+b)))
pr3 : n +N n <N a +N (b +N c)
pr3 = identityOfIndiscernablesRight _<N_ pr2'' (equalityCommutative (Semiring.+Associative Semiring a b c))
help15 : {a b c n : } (n<b+c : n <N b +N c) (n<a+b : n <N a +N b) (n <N a +N subtractionNResult.result (-N (inl n<b+c))) (subtractionNResult.result (-N (inl n<a+b)) +N c) <N n False
help15 {a} {b} {c} {n} n<b+c n<a+b pr1 pr2 = TotalOrder.irreflexive TotalOrder (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) (Semiring.+Associative Semiring n _ c)
p1' : (a +N b) +N c <N n +N n
p1' = identityOfIndiscernablesLeft _<N_ p1 (applyEquality (_+N c) (addMinus' (inl n<a+b)))
p2 : n +N n <N a +N (subtractionNResult.result (-N (inl n<b+c)) +N n)
p2 = identityOfIndiscernablesRight _<N_ (additionPreservesInequality n pr1) (equalityCommutative (Semiring.+Associative Semiring a _ n))
p2' : n +N n <N a +N (b +N c)
p2' = identityOfIndiscernablesRight _<N_ p2 (applyEquality (a +N_) (addMinus (inl n<b+c)))
p2'' : n +N n <N (a +N b) +N c
p2'' = identityOfIndiscernablesRight _<N_ p2' (Semiring.+Associative Semiring a b c)
help14 : {a b c n : } (n<b+c : n <N b +N c) (n<a+b : n <N a +N b) (pr1 : n <N a +N subtractionNResult.result (-N (inl n<b+c))) (pr2 : n <N subtractionNResult.result (-N (inl n<a+b)) +N c) subtractionNResult.result (-N (inl pr1)) subtractionNResult.result (-N (inl pr2))
help14 {a} {b} {c} {n} n<b+c n<a+b pr1 pr2 = equivalentSubtraction _ _ _ _ pr1 pr2 (transitivity (Semiring.+Associative Semiring n _ c) (transitivity (applyEquality (_+N c) (addMinus' (inl n<a+b))) (transitivity (equalityCommutative (Semiring.+Associative Semiring a b c)) (equalityCommutative p2))))
where
p1 : (a +N subtractionNResult.result (-N (inl n<b+c))) +N n a +N (subtractionNResult.result (-N (inl n<b+c)) +N n)
p1 = equalityCommutative (Semiring.+Associative Semiring a _ n)
p2 : (a +N subtractionNResult.result (-N (inl n<b+c))) +N n a +N (b +N c)
p2 = identityOfIndiscernablesRight _≡_ p1 (applyEquality (a +N_) (addMinus (inl n<b+c)))
help13 : {a b c n : } (n<b+c : n <N b +N c) (n<a+b : n <N a +N b) (n <N a +N subtractionNResult.result (-N (inl n<b+c))) (subtractionNResult.result (-N (inl n<a+b)) +N c n) False
help13 {a} {b} {c} {n} n<b+c n<a+b pr1 pr2 = TotalOrder.irreflexive TotalOrder (identityOfIndiscernablesRight _<N_ lemm1' lemm3)
where
lemm1 : n +N n <N a +N (subtractionNResult.result (-N (inl n<b+c)) +N n)
lemm1 = identityOfIndiscernablesRight _<N_ (additionPreservesInequality n pr1) (equalityCommutative (Semiring.+Associative Semiring a _ n))
lemm1' : n +N n <N a +N (b +N c)
lemm1' = identityOfIndiscernablesRight _<N_ lemm1 (applyEquality (a +N_) (addMinus (inl n<b+c)))
lemm2 : (n +N subtractionNResult.result (-N (inl n<a+b))) +N c n +N n
lemm2 = identityOfIndiscernablesLeft _≡_ (applyEquality (n +N_) pr2) (Semiring.+Associative Semiring n _ c)
lemm2' : (a +N b) +N c n +N n
lemm2' = identityOfIndiscernablesLeft _≡_ lemm2 (applyEquality (_+N c) (addMinus' (inl n<a+b)))
lemm3 : a +N (b +N c) n +N n
lemm3 rewrite Semiring.+Associative Semiring a b c = lemm2'
help12 : {a b c n : } (n<b+c : n <N b +N c) (n<a+b : n <N a +N b) (a +N subtractionNResult.result (-N (inl n<b+c))) n subtractionNResult.result (-N (inl n<a+b)) +N c <N n False
help12 {a} {b} {c} {n} n<b+c n<a+b pr1 pr2 = TotalOrder.irreflexive TotalOrder lemm4
where
pr : {a b c : } a +N (b +N c) b +N (a +N c)
pr {a} {b} {c} rewrite Semiring.+Associative Semiring a b c | Semiring.commutative Semiring a b | equalityCommutative (Semiring.+Associative Semiring b a c) = refl
lemm1 : (n +N subtractionNResult.result (-N (inl n<a+b))) +N c <N n +N n
lemm1 = identityOfIndiscernablesLeft _<N_ (additionPreservesInequalityOnLeft n pr2) (Semiring.+Associative Semiring n _ c)
lemm2 : (a +N b) +N c <N n +N n
lemm2 = identityOfIndiscernablesLeft _<N_ lemm1 (applyEquality (_+N c) (addMinus' (inl n<a+b)))
lemm1' : a +N (subtractionNResult.result (-N (inl n<b+c)) +N n) n +N n
lemm1' = identityOfIndiscernablesLeft _≡_ (applyEquality (_+N n) pr1) (equalityCommutative (Semiring.+Associative Semiring a _ n))
lemm2' : a +N (b +N c) n +N n
lemm2' = identityOfIndiscernablesLeft _≡_ lemm1' (applyEquality (a +N_) (addMinus (inl n<b+c)))
lemm3 : (a +N b) +N c n +N n
lemm3 rewrite equalityCommutative (Semiring.+Associative Semiring a b c) = lemm2'
lemm4 : (a +N b) +N c <N (a +N b) +N c
lemm4 = identityOfIndiscernablesRight _<N_ lemm2 (equalityCommutative lemm3)
help11 : {a b c n : } (a <N n) (b +N c n) (n<a+b : n <N a +N b) (n <N subtractionNResult.result (-N (inl n<a+b)) +N c) False
help11 {a} {b} {c} {n} a<n b+c=n n<a+b pr1 = TotalOrder.irreflexive TotalOrder (orderIsTransitive a<n lemm5)
where
pr : {a b c : } a +N (b +N c) b +N (a +N c)
pr {a} {b} {c} rewrite Semiring.+Associative Semiring a b c | Semiring.commutative Semiring a b | equalityCommutative (Semiring.+Associative Semiring b a c) = refl
lemm : n +N n <N (n +N subtractionNResult.result (-N (inl n<a+b))) +N c
lemm = identityOfIndiscernablesRight _<N_ (additionPreservesInequalityOnLeft n pr1) (Semiring.+Associative Semiring n _ c)
lemm2 : n +N n <N (a +N b) +N c
lemm2 = identityOfIndiscernablesRight _<N_ lemm (applyEquality (_+N c) (addMinus' (inl n<a+b)))
lemm3 : n +N n <N a +N (b +N c)
lemm3 = identityOfIndiscernablesRight _<N_ lemm2 (equalityCommutative (Semiring.+Associative Semiring a b c))
lemm4 : n +N n <N a +N n
lemm4 = identityOfIndiscernablesRight _<N_ lemm3 (applyEquality (a +N_) b+c=n)
lemm5 : n <N a
lemm5 = subtractionPreservesInequality n lemm4
help10 : {a b c n : } (n<b+c : n <N b +N c) (n<a+b : n <N a +N b) (a +N subtractionNResult.result (-N (inl n<b+c)) n) (n <N subtractionNResult.result (-N (inl n<a+b)) +N c) False
help10 {a} {b} {c} {n} n<b+c n<a+b pr1 pr2 = TotalOrder.irreflexive TotalOrder lemm6
where
pr : {a b c : } a +N (b +N c) b +N (a +N c)
pr {a} {b} {c} rewrite Semiring.+Associative Semiring a b c | Semiring.commutative Semiring a b | equalityCommutative (Semiring.+Associative Semiring b a c) = refl
lemm : a +N (n +N subtractionNResult.result (-N (inl n<b+c))) n +N n
lemm = identityOfIndiscernablesLeft _≡_ (applyEquality (n +N_) pr1) (pr {n} {a})
lemm2 : a +N (b +N c) n +N n
lemm2 = identityOfIndiscernablesLeft _≡_ lemm (applyEquality (a +N_) (addMinus' (inl n<b+c)))
lemm3 : n +N n <N (n +N subtractionNResult.result (-N (inl n<a+b))) +N c
lemm3 = identityOfIndiscernablesRight _<N_ (additionPreservesInequalityOnLeft n pr2) (Semiring.+Associative Semiring n _ c)
lemm4 : n +N n <N (a +N b) +N c
lemm4 = identityOfIndiscernablesRight _<N_ lemm3 (applyEquality (_+N c) (addMinus' (inl n<a+b)))
lemm5 : n +N n <N a +N (b +N c)
lemm5 = identityOfIndiscernablesRight _<N_ lemm4 (equalityCommutative (Semiring.+Associative Semiring a b c))
lemm6 : a +N (b +N c) <N a +N (b +N c)
lemm6 = identityOfIndiscernablesLeft _<N_ lemm5 (equalityCommutative lemm2)
help9 : {a n : } (a +N 0 n) (a <N n) False
help9 {a} {n} n=a+0 a<n rewrite Semiring.commutative Semiring a 0 | n=a+0 = TotalOrder.irreflexive TotalOrder a<n
help8 : {a n : } (n <N a +N 0) (a <N n) False
help8 {a} {n} n<a+0 a<n rewrite Semiring.commutative Semiring a 0 = TotalOrder.irreflexive TotalOrder (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 Semiring.commutative Semiring a 0 = canSubtractFromEqualityLeft {n} lem'
where
lem : n +N a (n +N subtractionNResult.result (-N (inl n<a+b))) +N c
lem rewrite addMinus' (inl n<a+b) | equalityCommutative (Semiring.+Associative Semiring a b c) | b+c=n = Semiring.commutative Semiring n a
lem' : n +N a n +N (subtractionNResult.result (-N (inl n<a+b)) +N c)
lem' = identityOfIndiscernablesRight _≡_ lem (equalityCommutative (Semiring.+Associative Semiring n _ c))
help5 : {a b c n : } (n<b+c : n <N b +N c) (n<a+b : n <N a +N b) a +N subtractionNResult.result (-N (inl n<b+c)) subtractionNResult.result (-N (inl n<a+b)) +N c
help5 {a} {b} {c} {n} n<b+c n<a+b = canSubtractFromEqualityLeft {n} lemma''
where
lemma : a +N (n +N subtractionNResult.result (-N (inl n<b+c))) (n +N subtractionNResult.result (-N (inl n<a+b))) +N c
lemma rewrite addMinus' (inl n<b+c) | addMinus' (inl n<a+b) = Semiring.+Associative Semiring a b c
lemma' : a +N (n +N subtractionNResult.result (-N (inl n<b+c))) n +N (subtractionNResult.result (-N (inl n<a+b)) +N c)
lemma' = identityOfIndiscernablesRight _≡_ lemma (equalityCommutative (Semiring.+Associative Semiring n _ c))
lemma'' : n +N (a +N subtractionNResult.result (-N (inl n<b+c))) n +N (subtractionNResult.result (-N (inl n<a+b)) +N c)
lemma'' = identityOfIndiscernablesLeft _≡_ lemma' (pr {a} {n} {subtractionNResult.result (-N (inl n<b+c))})
where
pr : {a b c : } a +N (b +N c) b +N (a +N c)
pr {a} {b} {c} rewrite Semiring.+Associative Semiring a b c | Semiring.commutative Semiring a b | equalityCommutative (Semiring.+Associative Semiring b a c) = refl
help4 : {a b c n : } (n<a+'b+c : n <N a +N (b +N c)) (n<a+b : n <N a +N b) (subtractionNResult.result (-N (inl n<a+'b+c)) subtractionNResult.result (-N (inl n<a+b)) +N c)
help4 {a} {b} {c} {n} n<a+'b+c n<a+b = canSubtractFromEqualityLeft lemma'
where
lemma : (n +N subtractionNResult.result (-N (inl n<a+'b+c))) (n +N subtractionNResult.result (-N (inl n<a+b))) +N c
lemma rewrite addMinus' (inl n<a+'b+c) | addMinus' (inl n<a+b) = Semiring.+Associative Semiring a b c
lemma' : n +N subtractionNResult.result (-N (inl n<a+'b+c)) n +N (subtractionNResult.result (-N (inl n<a+b)) +N c)
lemma' = identityOfIndiscernablesRight _≡_ lemma (equalityCommutative (Semiring.+Associative Semiring n (subtractionNResult.result (-N (inl n<a+b))) c))
help3 : {a b c n : } (a <N n) (b <N n) (c <N n) (a +N b <N n) (pr : n <N b +N c) a +N subtractionNResult.result (-N (inl pr)) n False
help3 {a} {b} {c} {n} a<n b<n c<n a+b<n n<b+c pr = TotalOrder.irreflexive TotalOrder (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 Semiring.+Associative Semiring a b c | Semiring.+Associative Semiring b a c = applyEquality (_+N c) (Semiring.commutative Semiring a b)
inter2 : n +N n a +N (b +N c)
inter2 = equalityCommutative (identityOfIndiscernablesLeft _≡_ inter (applyEquality (a +N_) (addMinus' (inl n<b+c))))
inter3 : n +N n <N n +N c
inter3 rewrite inter2 | Semiring.+Associative Semiring a b c = additionPreservesInequality c a+b<n
inter4 : (n +N n <N n +N c) n <N c
inter4 pr rewrite Semiring.commutative Semiring n c = subtractionPreservesInequality n pr
help2 : {a b c n : } (sn<b+c : succ n <N b +N c) (sn<a+b+c : succ n <N (a +N b) +N c) a +N subtractionNResult.result (-N (inl sn<b+c)) subtractionNResult.result (-N (inl sn<a+b+c))
help2 {a} {b} {c} {n} sn<b+c sn<a+b+c = res inter
where
inter : a +N (subtractionNResult.result (-N (inl sn<b+c)) +N succ n) subtractionNResult.result (-N (inl sn<a+b+c)) +N succ n
inter rewrite addMinus (inl sn<b+c) | addMinus (inl sn<a+b+c) = Semiring.+Associative Semiring a b c
res : (a +N (subtractionNResult.result (-N (inl sn<b+c)) +N succ n) subtractionNResult.result (-N (inl sn<a+b+c)) +N succ n) a +N subtractionNResult.result (-N (inl sn<b+c)) subtractionNResult.result (-N (inl sn<a+b+c))
res pr = canSubtractFromEqualityRight {_} {succ n} (identityOfIndiscernablesLeft _≡_ pr (Semiring.+Associative Semiring a (subtractionNResult.result (-N (inl sn<b+c))) (succ n)))
help1 : {a b c n : } (sn<b+c : succ n <N b +N c) (pr1 : succ n <N a +N subtractionNResult.result (-N (inl sn<b+c))) (a +N b <N succ n) (a <N succ n) (b <N succ n) (c <N succ n) False
help1 {a} {b} {c} {n} sn<b+c pr1 a+b<sn a<sn b<sn c<sn with -N (inl sn<b+c)
help1 {a} {b} {c} {n} sn<b+c pr1 a+b<sn a<sn b<sn c<sn | record { result = b+c-sn ; pr = Prb+c-sn } = ans
where
b+c-nNonzero : b+c-sn 0 False
b+c-nNonzero pr rewrite (equalityCommutative Prb+c-sn) | pr | Semiring.commutative Semiring n 0 = TotalOrder.irreflexive TotalOrder sn<b+c
2sn<a+b+c' : succ n +N succ n <N succ n +N (a +N b+c-sn)
2sn<a+b+c' = additionPreservesInequalityOnLeft (succ n) pr1
2sn<a+b+c'' : succ n +N succ n <N a +N (succ n +N b+c-sn)
2sn<a+b+c'' rewrite Semiring.+Associative Semiring a (succ n) b+c-sn | Semiring.commutative Semiring a (succ n) | equalityCommutative (Semiring.+Associative Semiring (succ n) a b+c-sn) = 2sn<a+b+c'
eep : succ n +N succ n <N a +N (b +N c)
eep rewrite equalityCommutative Prb+c-sn = 2sn<a+b+c''
eep2 : a +N (b +N c) <N succ n +N c
eep2 rewrite Semiring.+Associative Semiring a b c = additionPreservesInequality c a+b<sn
eep2' : a +N (b +N c) <N succ n +N succ n
eep2' = orderIsTransitive eep2 (additionPreservesInequalityOnLeft (succ n) c<sn)
ans : False
ans = TotalOrder.irreflexive TotalOrder (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 _ _ (Semiring.+Associative Semiring a b c)
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inl b+c<sn) | inl (inr sn<a+'b+c) = exFalso (false {succ n} a+b+c<sn sn<a+'b+c)
where
false : {x : } (a +N b) +N c <N succ n succ n <N a +N (b +N c) False
false pr1 pr2 rewrite equalityCommutative (Semiring.+Associative Semiring a b c) = TotalOrder.irreflexive TotalOrder (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 equalityCommutative (Semiring.+Associative Semiring a b c) | p2 = TotalOrder.irreflexive TotalOrder p1
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inr sn<b+c) with 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 | Semiring.+Associative Semiring a (succ n) result | Semiring.commutative Semiring a (succ n) | equalityCommutative (Semiring.+Associative Semiring (succ n) a result) = cannotAddAndEnlarge' pr1
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inr sn<b+c) | inl (inl _) | inl (inr x) with -N (inl sn<b+c)
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inr sn<b+c) | inl (inl a+'b+c<sn) | inl (inr x) | record { result = result ; pr = pr } = exFalso false
where
lemma : a +N (succ n +N result) a +N (b +N c)
lemma' : a +N (succ n +N result) <N succ n
lemma'' : succ n +N (a +N result) <N succ n
lemma'' = identityOfIndiscernablesLeft _<N_ lemma' (transitivity (Semiring.+Associative Semiring a (succ n) result) (transitivity (applyEquality (λ t t +N result) (Semiring.commutative Semiring a (succ n))) (equalityCommutative (Semiring.+Associative Semiring (succ n) a result))))
lemma = applyEquality (λ i a +N i) pr
lemma' rewrite lemma = a+'b+c<sn
false : False
false = cannotAddAndEnlarge' lemma''
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inr sn<b+c) | inl (inl _) | inr x with -N (inl sn<b+c)
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inr sn<b+c) | inl (inl a+'b+c<sn) | inr x | record { result = result ; pr = pr } = exFalso false
where
lemma : a +N (succ n +N result) <N succ n
lemma rewrite pr = a+'b+c<sn
lemma' : succ n +N (a +N result) <N succ n
lemma' = identityOfIndiscernablesLeft _<N_ lemma (transitivity (Semiring.+Associative Semiring a (succ n) result) (transitivity (applyEquality (λ t t +N result) (Semiring.commutative Semiring a (succ n))) (equalityCommutative (Semiring.+Associative Semiring (succ n) a result))))
false : False
false = cannotAddAndEnlarge' lemma'
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inr sn<b+c) | inl (inr x) = equalityZn _ _ (exFalso (false {succ n} a+b+c<sn x))
where
false : {x : } (a +N b) +N c <N succ n succ n <N a +N (b +N c) False
false pr1 pr2 rewrite equalityCommutative (Semiring.+Associative Semiring a b c) = TotalOrder.irreflexive TotalOrder (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 | equalityCommutative (Semiring.+Associative Semiring a b c) = TotalOrder.irreflexive TotalOrder pr1
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inr sn=b+c = exFalso (false a+b+c<sn sn=b+c)
where
false : (a +N b) +N c <N succ n b +N c succ n False
false pr1 pr2 rewrite equalityCommutative (Semiring.+Associative Semiring a b c) | pr2 | Semiring.commutative Semiring a (succ n) = cannotAddAndEnlarge' a+b+c<sn
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) with 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 equalityCommutative (Semiring.+Associative Semiring a b c) = TotalOrder.irreflexive TotalOrder (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 equalityCommutative (Semiring.+Associative Semiring a b c) = Semiring.commutative Semiring (succ n) _
ans : subtractionNResult.result (-N (inl sn<a+'b+c)) subtractionNResult.result (-N (inl sn<a+b+c))
ans = equivalentSubtraction _ _ _ _ sn<a+'b+c sn<a+b+c lemma
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inl (inl b+c<sn) | inr x = exFalso (false sn<a+b+c x)
where
false : succ n <N (a +N b) +N c (a +N (b +N c)) succ n False
false pr1 pr2 rewrite equalityCommutative (Semiring.+Associative Semiring a b c) | pr2 = TotalOrder.irreflexive TotalOrder sn<a+b+c
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inl (inr sn<b+c) with 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 equalityCommutative (Semiring.+Associative Semiring a b c) = TotalOrder.irreflexive TotalOrder (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 equalityCommutative (Semiring.+Associative Semiring a b c) | pr2 = TotalOrder.irreflexive TotalOrder pr1
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inr sn=b+c with 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 equalityCommutative (Semiring.+Associative Semiring a b c) = TotalOrder.irreflexive TotalOrder (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 Semiring.commutative Semiring a 0 | equalityCommutative (Semiring.+Associative Semiring a b c) | b+c=sn | Semiring.commutative Semiring (succ n) result = equalityCommutative (canSubtractFromEqualityRight pr1)
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inr sn=b+c | inl (inr _) | inl (inr x) = exFalso (false b a+b<sn x)
where
false : (b : ) a +N b <N succ n succ n <N a +N 0 False
false zero pr1 pr2 rewrite Semiring.commutative Semiring a 0 = TotalOrder.irreflexive TotalOrder (orderIsTransitive pr1 pr2)
false (succ b) pr1 pr2 rewrite Semiring.commutative Semiring a 0 = TotalOrder.irreflexive TotalOrder (orderIsTransitive pr2 (orderIsTransitive (le b (Semiring.commutative Semiring (succ b) a)) pr1))
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inr sn=b+c | inl (inr _) | inr x = exFalso (false b a+b<sn x)
where
false : (b : ) a +N b <N succ n a +N 0 succ n False
false zero pr1 pr2 rewrite pr2 = TotalOrder.irreflexive TotalOrder pr1
false (succ b) pr1 pr2 rewrite Semiring.commutative Semiring a 0 | pr2 = cannotAddAndEnlarge' pr1
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inr sn=b+c | inr x = exFalso (false sn<a+b+c x)
where
false : succ n <N (a +N b) +N c a +N (b +N c) succ n False
false pr1 pr2 rewrite equalityCommutative (Semiring.+Associative Semiring a b c) | pr2 = TotalOrder.irreflexive TotalOrder pr1
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inr sn=a+b+c with 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 equalityCommutative (Semiring.+Associative Semiring a b c) | pr1 = TotalOrder.irreflexive TotalOrder pr2
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inr sn=a+b+c | inl (inl b+c<sn) | inl (inr sn<a+'b+c) = exFalso (false sn=a+b+c sn<a+'b+c)
where
false : (a +N b) +N c succ n succ n <N a +N (b +N c) False
false pr1 pr2 rewrite equalityCommutative (Semiring.+Associative Semiring a b c) | pr1 = TotalOrder.irreflexive TotalOrder pr2
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inr sn=a+b+c | inl (inl b+c<sn) | inr _ = equalityZn _ _ refl
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inr sn=a+b+c | inl (inr sn<b+c) = exFalso (false a sn=a+b+c sn<b+c)
where
false : (a : ) (a +N b) +N c succ n succ n <N b +N c False
false zero pr1 pr2 rewrite equalityCommutative (Semiring.+Associative Semiring a b c) | equalityCommutative pr1 = TotalOrder.irreflexive TotalOrder pr2
false (succ a) pr1 pr2 rewrite equalityCommutative (Semiring.+Associative Semiring a b c) | equalityCommutative pr1 = TotalOrder.irreflexive TotalOrder (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 equalityCommutative (Semiring.+Associative Semiring a b c) | pr2 = refl
a=0 (succ a) pr1 pr2 rewrite equalityCommutative (Semiring.+Associative Semiring a b c) | pr2 = canSubtractFromEqualityRight pr1
ans : a +N 0 0
ans rewrite Semiring.commutative Semiring a 0 = a=0 a sn=a+b+c b+c=sn
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inr sn=a+b+c | inr b+c=sn | inl (inr sn<a+0) = exFalso (false sn<a+0 sn=a+b+c b+c=sn)
where
false : succ n <N a +N 0 (a +N b) +N c succ n b +N c succ n False
false pr1 pr2 pr3 rewrite equalityCommutative (Semiring.+Associative Semiring a b c) | pr3 | Semiring.commutative Semiring a 0 = zeroNeverGreater {succ n} (identityOfIndiscernablesRight _<N_ pr1 (a=0 a pr2))
where
a=0 : (a : ) (a +N succ n succ n) a 0
a=0 zero pr = refl
a=0 (succ a) pr = canSubtractFromEqualityRight pr
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inr sn=a+b+c | inr b+c=sn | inr sn=a+0 = exFalso (false sn=a+b+c b+c=sn sn=a+0)
where
false : (a +N b) +N c succ n b +N c succ n a +N 0 succ n False
false pr1 pr2 pr3 rewrite equalityCommutative (Semiring.+Associative Semiring a b c) | pr2 | equalityCommutative pr3 | Semiring.commutative Semiring a 0 = naughtE (identityOfIndiscernablesLeft _≡_ pr3 (a=0 a pr1))
where
a=0 : (a : ) (a +N a a) a 0
a=0 zero pr = refl
a=0 (succ a) pr = exFalso (naughtE {a} (equalityCommutative (canSubtractFromEqualityRight pr)))
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) with 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 Semiring.+Associative Semiring 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 Semiring.+Associative Semiring a b c | equalityCommutative pr2 = cannotAddAndEnlarge' pr1
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inl (inr sn<b+c) with 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 (TotalOrder.irreflexive TotalOrder (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 | Semiring.commutative Semiring b c = cannotAddAndEnlarge' b+c<sn
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inr c=sn | inl (inr sn<b+c) with 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 (TotalOrder.irreflexive TotalOrder 0<a)
subLess {succ a} {b} _ a<b with -N (inl a<b)
... | record { result = b-a ; pr = pr } = record { x = a ; proof = pr }
inverseZn : {n : } {pr : 0 <N n} (a : n n pr) Sg (n n pr) (λ i i +n a record { x = 0 ; xLess = pr } )
inverseZn {zero} {()}
inverseZn {succ n} {0<n} record { x = zero ; xLess = zeroLess } = record { x = zero ; xLess = zeroLess } , plusZnIdentityLeft _
inverseZn {succ n} {0<n} record { x = (succ a) ; xLess = aLess } = ans , pr
where
ans = record { x = subtractionNResult.result (-N (inl (canRemoveSuccFrom<N aLess))) ; xLess = subLess (succIsPositive a) aLess }
pr : ans +n record { x = (succ a) ; xLess = aLess } record { x = 0 ; xLess = 0<n }
pr with orderIsTotal (subtractionNResult.result (-N (inl (canRemoveSuccFrom<N aLess))) +N (succ a)) (succ n)
... | inl (inl x) = exFalso f
where
h : subtractionNResult.result (-N (inl (canRemoveSuccFrom<N aLess))) +N succ a succ n
h with -N (inl (canRemoveSuccFrom<N aLess))
h | record { result = result ; pr = pr } rewrite equalityCommutative pr = Semiring.commutative Semiring result (succ a)
f : False
f = TotalOrder.irreflexive TotalOrder (identityOfIndiscernablesLeft _<N_ x h)
... | inl (inr x) = exFalso f
where
h : subtractionNResult.result (-N (inl (canRemoveSuccFrom<N aLess))) +N succ a succ n
h = addMinus {succ a} {succ n} (inl aLess)
f : False
f rewrite h = TotalOrder.irreflexive TotalOrder x
... | (inr n-a+sa=sn) = equalityZn _ _ refl
nGroup : (n : ) (pr : 0 <N n) Group (reflSetoid (n n pr)) _+n_
nGroup n pr = record { 0G = record { x = 0 ; xLess = pr } ; inverse = λ a underlying (inverseZn a) ; +Associative = λ {a} {b} {c} plusZnAssociative a b c ; identRight = λ {a} plusZnIdentityRight a ; identLeft = λ {a} plusZnIdentityLeft a ; invLeft = λ {a} helpInvLeft a ; invRight = λ {a} helpInvRight a ; +WellDefined = reflGroupWellDefined }
where
helpInvLeft : (a : n n pr) underlying (inverseZn a) +n a record { x = 0 ; xLess = pr }
helpInvLeft a with inverseZn a
... | -a , pr-a = pr-a
helpInvRight : (a : n n pr) a +n underlying (inverseZn a) record { x = 0 ; xLess = pr }
helpInvRight a rewrite plusZnCommutative a (underlying (inverseZn a)) = helpInvLeft a
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

View File

@@ -1,705 +0,0 @@
{-# OPTIONS --safe --warning=error #-}
-- These are explicitly with-K, because we currently encode an element of Zn as
-- a natural together with a proof that it is small.
open import LogicalFormulae
open import Groups.Definition
open import Groups.Groups
open import Numbers.Naturals.Naturals
open import Numbers.Naturals.Addition -- TODO remove this dependency
open import Numbers.Primes.PrimeNumbers
open import Setoids.Setoids
open import Sets.FinSet
open import Sets.FinSetWithK
open import Functions
open import Numbers.Naturals.WithK
open import Semirings.Definition
open import Orders
module Numbers.Modulo.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 Semiring.+Associative Semiring (succ x) a b | Semiring.commutative Semiring x a | equalityCommutative (Semiring.+Associative Semiring (succ a) x b) | Semiring.commutative Semiring a (x +N b) | Semiring.commutative Semiring (succ (x +N b)) a = canSubtractFromEqualityLeft {a} proof
_+n_ : {n : } {pr : 0 <N n} n n pr n n pr n n pr
_+n_ {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 _ _ (Semiring.commutative Semiring a 0)
plusZnIdentityRight {succ x} {_} record { x = a ; xLess = aLess } | inl (inr sx<a) = exFalso (f aLess sx<a)
where
f : (aL : a <N succ x) (sx<a : succ x <N a +N 0) False
f aL sx<a rewrite Semiring.commutative Semiring a 0 = 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 Semiring.commutative Semiring a 0 | a=sx = TotalOrder.irreflexive TotalOrder 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 (TotalOrder.irreflexive TotalOrder (TotalOrder.<Transitive TotalOrder succn<x xLess))
plusZnIdentityLeft {succ n} {pr} record { x = x ; xLess = xLess } | inr x=succn rewrite x=succn = exFalso (TotalOrder.irreflexive TotalOrder xLess)
subLemma : {a b c : } (pr1 : a <N b) (pr2 : a <N c) (pr : b c) (subtractionNResult.result (-N (inl pr1))) (subtractionNResult.result (-N (inl pr2)))
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 _ _ (Semiring.commutative Semiring a b)
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inl (inl a+b<sn) | inl (inr sn<b+a) = exFalso (f a+b<sn sn<b+a)
where
f : (a +N b) <N succ n succ n <N b +N a False
f pr1 pr2 rewrite Semiring.commutative Semiring b a = TotalOrder.irreflexive TotalOrder (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 Semiring.commutative Semiring b a | eq = TotalOrder.irreflexive TotalOrder pr1
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inl (inr sn<a+b) with 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 Semiring.commutative Semiring b a = TotalOrder.irreflexive TotalOrder (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 (Semiring.commutative Semiring a b))
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inl (inr sn<a+b) | inr sn=b+a = exFalso (f sn<a+b sn=b+a)
where
f : succ n <N a +N b b +N a succ n False
f pr eq rewrite Semiring.commutative Semiring b a | eq = TotalOrder.irreflexive TotalOrder pr
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inr sn=a+b with 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 Semiring.commutative Semiring b a | sn=a+b = TotalOrder.irreflexive TotalOrder b+a<sn
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inr sn=a+b | inl (inr sn<b+a) = exFalso f
where
f : False
f rewrite Semiring.commutative Semiring b a | sn=a+b = TotalOrder.irreflexive TotalOrder sn<b+a
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inr sn=a+b | inr sn=b+a = equalityZn _ _ refl
help30 : {a b c n : } (c <N n) (a +N b n) (n<b+c : n <N b +N c) (n <N a +N subtractionNResult.result (-N (inl n<b+c))) False
help30 {a} {b} {c} {n} c<n a+b=n n<b+c x = TotalOrder.irreflexive TotalOrder (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) (equalityCommutative (Semiring.+Associative Semiring a _ n))
pr2 : n +N n <N a +N (b +N c)
pr2 = identityOfIndiscernablesRight _<N_ pr (applyEquality (a +N_) (addMinus (inl n<b+c)))
pr3 : n +N n <N (a +N b) +N c
pr3 rewrite equalityCommutative (Semiring.+Associative Semiring a b c) = pr2
pr4 : n +N n <N c +N n
pr4 rewrite Semiring.commutative Semiring c n = identityOfIndiscernablesRight _<N_ pr3 (applyEquality (_+N c) a+b=n)
pr5 : n <N c
pr5 = subtractionPreservesInequality n pr4
help31 : {a b c n : } (a +N b n) (n<b+c : n <N b +N c) (a +N subtractionNResult.result (-N (inl n<b+c))) c
help31 {a} {b} {c} {n} a+b=n n<b+c = canSubtractFromEqualityLeft pr2
where
pr1 : a +N (n +N subtractionNResult.result (-N (inl n<b+c))) n +N c
pr1 rewrite addMinus' (inl n<b+c) | Semiring.+Associative Semiring a b c | a+b=n = refl
pr2 : n +N (a +N subtractionNResult.result (-N (inl n<b+c))) n +N c
pr2 = identityOfIndiscernablesLeft _≡_ pr1 (lemm a n (subtractionNResult.result (-N (inl n<b+c))))
where
lemm : (a b c : ) a +N (b +N c) b +N (a +N c)
lemm a b c rewrite Semiring.+Associative Semiring a b c | Semiring.commutative Semiring a b | equalityCommutative (Semiring.+Associative Semiring b a c) = refl
help7 : {a b c n : } b +N c n a <N n (n<a+b : n <N a +N b) (subtractionNResult.result (-N (inl n<a+b)) +N c n) False
help7 {a} {b} {c} {n} b+c=n a<n n<a+b x = TotalOrder.irreflexive TotalOrder (identityOfIndiscernablesLeft _<N_ a<n pr5)
where
pr : (n +N subtractionNResult.result (-N (inl n<a+b))) +N c n +N n
pr = identityOfIndiscernablesLeft _≡_ (applyEquality (n +N_) x) (Semiring.+Associative Semiring n _ c)
pr2 : (a +N b) +N c n +N n
pr2 = identityOfIndiscernablesLeft _≡_ pr (applyEquality (_+N c) (addMinus' (inl n<a+b)))
pr3 : a +N (b +N c) n +N n
pr3 rewrite Semiring.+Associative Semiring a b c = pr2
pr4 : a +N n n +N n
pr4 = identityOfIndiscernablesLeft _≡_ pr3 (applyEquality (a +N_) b+c=n)
pr5 : a n
pr5 = canSubtractFromEqualityRight pr4
help29 : {a b c n : } (c <N n) (n<b+c : n <N b +N c) (a +N subtractionNResult.result (-N (inl n<b+c))) n a +N b n False
help29 {a} {b} {c} {n} c<n n<b+c pr a+b=n = TotalOrder.irreflexive TotalOrder (identityOfIndiscernablesLeft _<N_ c<n p4)
where
p1 : a +N (subtractionNResult.result (-N (inl n<b+c)) +N n) n +N n
p1 = identityOfIndiscernablesLeft _≡_ (applyEquality (_+N n) pr) (equalityCommutative (Semiring.+Associative Semiring a _ n))
p2 : (a +N b) +N c n +N n
p2 rewrite equalityCommutative (Semiring.+Associative Semiring a b c) = identityOfIndiscernablesLeft _≡_ p1 (applyEquality (a +N_) (addMinus (inl n<b+c)))
p3 : n +N c n +N n
p3 = identityOfIndiscernablesLeft _≡_ p2 (applyEquality (_+N c) a+b=n)
p4 : c n
p4 = canSubtractFromEqualityLeft p3
help28 : {a b c n : } (n<a+'b+c : n <N a +N (b +N c)) (a +N b n) subtractionNResult.result (-N (inl n<a+'b+c)) c
help28 {a} {b} {c} {n} pr a+b=n = canSubtractFromEqualityLeft p2
where
p1 : a +N (b +N c) n +N c
p1 rewrite Semiring.+Associative Semiring a b c | a+b=n = refl
p2 : n +N subtractionNResult.result (-N (inl pr)) n +N c
p2 = identityOfIndiscernablesLeft _≡_ p1 (equalityCommutative (addMinus' (inl pr)))
help27 : {a b c n : } (a +N b succ n) (a +N (b +N c) <N succ n) a +N (b +N c) c
help27 {a} {b} {zero} {n} a+b=sn a+b+c<sn rewrite Semiring.commutative Semiring b 0 | a+b=sn = exFalso (TotalOrder.irreflexive TotalOrder a+b+c<sn)
help27 {a} {b} {succ c} {n} a+b=sn a+b+c<sn rewrite Semiring.+Associative Semiring a b (succ c) | a+b=sn = exFalso (cannotAddAndEnlarge' a+b+c<sn)
help26 : {a b c n : } (a +N b n) (a +N (b +N c) n) 0 c
help26 {a} {b} {c} {n} a+b=n a+b+c=n rewrite Semiring.+Associative Semiring a b c | a+b=n | Semiring.commutative Semiring n c = equalityCommutative (canSubtractFromEqualityRight a+b+c=n)
help25 : {a b c n : } (a +N b n) (b +N c n) (a +N 0 c)
help25 {a} {b} {c} {n} a+b=n b+c=n rewrite Semiring.commutative Semiring a 0 | Semiring.commutative Semiring a b | equalityCommutative a+b=n = equalityCommutative (canSubtractFromEqualityLeft b+c=n)
help24 : {a n : } (a <N n) (n <N a +N 0) False
help24 {a} {n} a<n n<a+0 rewrite Semiring.commutative Semiring a 0 = TotalOrder.irreflexive TotalOrder (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 Semiring.commutative Semiring a 0 | a+0=n = TotalOrder.irreflexive TotalOrder a<n
help22 : {a b c n : } (a +N b n) (c n) (n<b+c : n <N b +N c) (n <N a +N subtractionNResult.result (-N (inl n<b+c))) False
help22 {a} {b} {c} {.c} a+b=n refl n<b+c pr = TotalOrder.irreflexive TotalOrder (identityOfIndiscernablesRight _<N_ pr4 a+b=n)
where
pr1 : c +N c <N a +N (subtractionNResult.result (-N (inl n<b+c)) +N c)
pr1 = identityOfIndiscernablesRight _<N_ (additionPreservesInequality c pr) (equalityCommutative (Semiring.+Associative Semiring a _ c))
pr2 : c +N c <N a +N (b +N c)
pr2 = identityOfIndiscernablesRight _<N_ pr1 (applyEquality (a +N_) (addMinus (inl n<b+c)))
pr3 : c +N c <N (a +N b) +N c
pr3 = identityOfIndiscernablesRight _<N_ pr2 (Semiring.+Associative Semiring a b c)
pr4 : c <N a +N b
pr4 = subtractionPreservesInequality c pr3
help21 : {a b c n : } (a +N b n) (b +N c n) (c n) (a <N n) False
help21 {a} {b} {c} {.c} a+b=n b+c=n refl a<n = TotalOrder.irreflexive TotalOrder (identityOfIndiscernablesLeft _<N_ a<n a=c)
where
b=0 : b 0
a=c : a c
a=c = identityOfIndiscernablesLeft _≡_ a+b=n lemm
where
lemm : a +N b a
lemm rewrite b=0 | Semiring.commutative Semiring a 0 = refl
b=0' : (b c : ) (b +N c c) b 0
b=0' zero c b+c=c = refl
b=0' (succ b) zero b+c=c = exFalso (naughtE (equalityCommutative b+c=c))
b=0' (succ b) (succ c) b+c=c = b=0' (succ b) c bl2
where
bl2 : succ b +N c c
bl2 rewrite Semiring.commutative Semiring b c | Semiring.commutative Semiring (succ c) b = succInjective b+c=c
b=0 = b=0' b c b+c=n
help20 : {a b c n : } (c n) (a +N b n) (n<b+c : n <N b +N c) (a +N subtractionNResult.result (-N (inl n<b+c)) <N n) False
help20 {a} {b} {c} {n} c=n a+b=n n<b+c pr = TotalOrder.irreflexive TotalOrder (identityOfIndiscernablesLeft _<N_ pr5 c=n)
where
pr1 : a +N (subtractionNResult.result (-N (inl n<b+c)) +N n) <N n +N n
pr1 = identityOfIndiscernablesLeft _<N_ (additionPreservesInequality n pr) (equalityCommutative (Semiring.+Associative Semiring a _ n))
pr2 : a +N (b +N c) <N n +N n
pr2 = identityOfIndiscernablesLeft _<N_ pr1 (applyEquality (a +N_) (addMinus (inl n<b+c)))
pr3 : (a +N b) +N c <N n +N n
pr3 rewrite equalityCommutative (Semiring.+Associative Semiring a b c) = pr2
pr4 : c +N n <N n +N n
pr4 = identityOfIndiscernablesLeft _<N_ pr3 (transitivity (applyEquality (_+N c) a+b=n) (Semiring.commutative Semiring n c))
pr5 : c <N n
pr5 = subtractionPreservesInequality n pr4
help19 : {a b c n : } (b+c<n : b +N c <N n) (n<a+b : n <N a +N b) (a <N n) (subtractionNResult.result (-N (inl n<a+b)) +N c n) False
help19 {a} {b} {c} {n} b+c<n n<a+b a<n pr = TotalOrder.irreflexive TotalOrder (identityOfIndiscernablesLeft _<N_ r q')
where
p : (n +N subtractionNResult.result (-N (inl n<a+b))) +N c n +N n
p = identityOfIndiscernablesLeft _≡_ (applyEquality (n +N_ ) pr) (Semiring.+Associative Semiring n _ c)
q : (a +N b) +N c n +N n
q = identityOfIndiscernablesLeft _≡_ p (applyEquality (_+N c) (addMinus' (inl n<a+b)))
q' : a +N (b +N c) n +N n
q' = identityOfIndiscernablesLeft _≡_ q (equalityCommutative (Semiring.+Associative Semiring a b c))
r : a +N (b +N c) <N n +N n
r = addStrongInequalities a<n b+c<n
help18 : {a b c n : } (b+c<n : b +N c <N n) (n<a+b : n <N a +N b) (a <N n) (n <N subtractionNResult.result (-N (inl n<a+b)) +N c) False
help18 {a} {b} {c} {n} b+c<n n<a+b a<n pr = TotalOrder.irreflexive TotalOrder (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) (Semiring.+Associative Semiring n _ c)
p' : n +N n <N (a +N b) +N c
p' = identityOfIndiscernablesRight _<N_ p (applyEquality (_+N c) (addMinus' (inl n<a+b)))
p2 : n +N n <N a +N (b +N c)
p2 = identityOfIndiscernablesRight _<N_ p' (equalityCommutative (Semiring.+Associative Semiring a b c))
p3 : n +N n <N a +N n
p3 = 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 = TotalOrder.irreflexive TotalOrder (identityOfIndiscernablesLeft _<N_ pr1'' pr3)
where
pr1' : a +N (subtractionNResult.result (-N (inl n<b+c)) +N n) <N n +N n
pr1' = identityOfIndiscernablesLeft _<N_ (additionPreservesInequality n p1) (equalityCommutative (Semiring.+Associative Semiring a _ n))
pr1'' : a +N (b +N c) <N n +N n
pr1'' = identityOfIndiscernablesLeft _<N_ pr1' (applyEquality (a +N_) (addMinus (inl n<b+c)))
pr2' : (n +N subtractionNResult.result (-N (inl n<a+b))) +N c n +N n
pr2' = identityOfIndiscernablesLeft _≡_ (applyEquality (n +N_) p2) (Semiring.+Associative Semiring n _ c)
pr2'' : (a +N b) +N c n +N n
pr2'' = identityOfIndiscernablesLeft _≡_ pr2' (applyEquality (_+N c) (addMinus' (inl n<a+b)))
pr3 : a +N (b +N c) n +N n
pr3 = identityOfIndiscernablesLeft _≡_ pr2'' (equalityCommutative (Semiring.+Associative Semiring a b c))
help16 : {a b c n : } (n<b+c : n <N b +N c) (n<a+b : n <N a +N b) (a +N subtractionNResult.result (-N (inl n<b+c))) <N n (pr : n <N subtractionNResult.result (-N (inl n<a+b)) +N c) a +N subtractionNResult.result (-N (inl n<b+c)) subtractionNResult.result (-N (inl pr))
help16 {a} {b} {c} {n} n<b+c n<a+b pr1 pr2 = exFalso (TotalOrder.irreflexive TotalOrder (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) (equalityCommutative (Semiring.+Associative Semiring a _ n))
pr1'' : a +N (b +N c) <N n +N n
pr1'' = identityOfIndiscernablesLeft _<N_ pr1' (applyEquality (a +N_) (addMinus (inl n<b+c)))
pr2' : n +N n <N (n +N subtractionNResult.result (-N (inl n<a+b))) +N c
pr2' = identityOfIndiscernablesRight _<N_ (additionPreservesInequalityOnLeft n pr2) (Semiring.+Associative Semiring n _ c)
pr2'' : n +N n <N (a +N b) +N c
pr2'' = identityOfIndiscernablesRight _<N_ pr2' (applyEquality (_+N c) (addMinus' (inl n<a+b)))
pr3 : n +N n <N a +N (b +N c)
pr3 = identityOfIndiscernablesRight _<N_ pr2'' (equalityCommutative (Semiring.+Associative Semiring a b c))
help15 : {a b c n : } (n<b+c : n <N b +N c) (n<a+b : n <N a +N b) (n <N a +N subtractionNResult.result (-N (inl n<b+c))) (subtractionNResult.result (-N (inl n<a+b)) +N c) <N n False
help15 {a} {b} {c} {n} n<b+c n<a+b pr1 pr2 = TotalOrder.irreflexive TotalOrder (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) (Semiring.+Associative Semiring n _ c)
p1' : (a +N b) +N c <N n +N n
p1' = identityOfIndiscernablesLeft _<N_ p1 (applyEquality (_+N c) (addMinus' (inl n<a+b)))
p2 : n +N n <N a +N (subtractionNResult.result (-N (inl n<b+c)) +N n)
p2 = identityOfIndiscernablesRight _<N_ (additionPreservesInequality n pr1) (equalityCommutative (Semiring.+Associative Semiring a _ n))
p2' : n +N n <N a +N (b +N c)
p2' = identityOfIndiscernablesRight _<N_ p2 (applyEquality (a +N_) (addMinus (inl n<b+c)))
p2'' : n +N n <N (a +N b) +N c
p2'' = identityOfIndiscernablesRight _<N_ p2' (Semiring.+Associative Semiring a b c)
help14 : {a b c n : } (n<b+c : n <N b +N c) (n<a+b : n <N a +N b) (pr1 : n <N a +N subtractionNResult.result (-N (inl n<b+c))) (pr2 : n <N subtractionNResult.result (-N (inl n<a+b)) +N c) subtractionNResult.result (-N (inl pr1)) subtractionNResult.result (-N (inl pr2))
help14 {a} {b} {c} {n} n<b+c n<a+b pr1 pr2 = equivalentSubtraction _ _ _ _ pr1 pr2 (transitivity (Semiring.+Associative Semiring n _ c) (transitivity (applyEquality (_+N c) (addMinus' (inl n<a+b))) (transitivity (equalityCommutative (Semiring.+Associative Semiring a b c)) (equalityCommutative p2))))
where
p1 : (a +N subtractionNResult.result (-N (inl n<b+c))) +N n a +N (subtractionNResult.result (-N (inl n<b+c)) +N n)
p1 = equalityCommutative (Semiring.+Associative Semiring a _ n)
p2 : (a +N subtractionNResult.result (-N (inl n<b+c))) +N n a +N (b +N c)
p2 = identityOfIndiscernablesRight _≡_ p1 (applyEquality (a +N_) (addMinus (inl n<b+c)))
help13 : {a b c n : } (n<b+c : n <N b +N c) (n<a+b : n <N a +N b) (n <N a +N subtractionNResult.result (-N (inl n<b+c))) (subtractionNResult.result (-N (inl n<a+b)) +N c n) False
help13 {a} {b} {c} {n} n<b+c n<a+b pr1 pr2 = TotalOrder.irreflexive TotalOrder (identityOfIndiscernablesRight _<N_ lemm1' lemm3)
where
lemm1 : n +N n <N a +N (subtractionNResult.result (-N (inl n<b+c)) +N n)
lemm1 = identityOfIndiscernablesRight _<N_ (additionPreservesInequality n pr1) (equalityCommutative (Semiring.+Associative Semiring a _ n))
lemm1' : n +N n <N a +N (b +N c)
lemm1' = identityOfIndiscernablesRight _<N_ lemm1 (applyEquality (a +N_) (addMinus (inl n<b+c)))
lemm2 : (n +N subtractionNResult.result (-N (inl n<a+b))) +N c n +N n
lemm2 = identityOfIndiscernablesLeft _≡_ (applyEquality (n +N_) pr2) (Semiring.+Associative Semiring n _ c)
lemm2' : (a +N b) +N c n +N n
lemm2' = identityOfIndiscernablesLeft _≡_ lemm2 (applyEquality (_+N c) (addMinus' (inl n<a+b)))
lemm3 : a +N (b +N c) n +N n
lemm3 rewrite Semiring.+Associative Semiring a b c = lemm2'
help12 : {a b c n : } (n<b+c : n <N b +N c) (n<a+b : n <N a +N b) (a +N subtractionNResult.result (-N (inl n<b+c))) n subtractionNResult.result (-N (inl n<a+b)) +N c <N n False
help12 {a} {b} {c} {n} n<b+c n<a+b pr1 pr2 = TotalOrder.irreflexive TotalOrder lemm4
where
pr : {a b c : } a +N (b +N c) b +N (a +N c)
pr {a} {b} {c} rewrite Semiring.+Associative Semiring a b c | Semiring.commutative Semiring a b | equalityCommutative (Semiring.+Associative Semiring b a c) = refl
lemm1 : (n +N subtractionNResult.result (-N (inl n<a+b))) +N c <N n +N n
lemm1 = identityOfIndiscernablesLeft _<N_ (additionPreservesInequalityOnLeft n pr2) (Semiring.+Associative Semiring n _ c)
lemm2 : (a +N b) +N c <N n +N n
lemm2 = identityOfIndiscernablesLeft _<N_ lemm1 (applyEquality (_+N c) (addMinus' (inl n<a+b)))
lemm1' : a +N (subtractionNResult.result (-N (inl n<b+c)) +N n) n +N n
lemm1' = identityOfIndiscernablesLeft _≡_ (applyEquality (_+N n) pr1) (equalityCommutative (Semiring.+Associative Semiring a _ n))
lemm2' : a +N (b +N c) n +N n
lemm2' = identityOfIndiscernablesLeft _≡_ lemm1' (applyEquality (a +N_) (addMinus (inl n<b+c)))
lemm3 : (a +N b) +N c n +N n
lemm3 rewrite equalityCommutative (Semiring.+Associative Semiring a b c) = lemm2'
lemm4 : (a +N b) +N c <N (a +N b) +N c
lemm4 = identityOfIndiscernablesRight _<N_ lemm2 (equalityCommutative lemm3)
help11 : {a b c n : } (a <N n) (b +N c n) (n<a+b : n <N a +N b) (n <N subtractionNResult.result (-N (inl n<a+b)) +N c) False
help11 {a} {b} {c} {n} a<n b+c=n n<a+b pr1 = TotalOrder.irreflexive TotalOrder (orderIsTransitive a<n lemm5)
where
pr : {a b c : } a +N (b +N c) b +N (a +N c)
pr {a} {b} {c} rewrite Semiring.+Associative Semiring a b c | Semiring.commutative Semiring a b | equalityCommutative (Semiring.+Associative Semiring b a c) = refl
lemm : n +N n <N (n +N subtractionNResult.result (-N (inl n<a+b))) +N c
lemm = identityOfIndiscernablesRight _<N_ (additionPreservesInequalityOnLeft n pr1) (Semiring.+Associative Semiring n _ c)
lemm2 : n +N n <N (a +N b) +N c
lemm2 = identityOfIndiscernablesRight _<N_ lemm (applyEquality (_+N c) (addMinus' (inl n<a+b)))
lemm3 : n +N n <N a +N (b +N c)
lemm3 = identityOfIndiscernablesRight _<N_ lemm2 (equalityCommutative (Semiring.+Associative Semiring a b c))
lemm4 : n +N n <N a +N n
lemm4 = identityOfIndiscernablesRight _<N_ lemm3 (applyEquality (a +N_) b+c=n)
lemm5 : n <N a
lemm5 = subtractionPreservesInequality n lemm4
help10 : {a b c n : } (n<b+c : n <N b +N c) (n<a+b : n <N a +N b) (a +N subtractionNResult.result (-N (inl n<b+c)) n) (n <N subtractionNResult.result (-N (inl n<a+b)) +N c) False
help10 {a} {b} {c} {n} n<b+c n<a+b pr1 pr2 = TotalOrder.irreflexive TotalOrder lemm6
where
pr : {a b c : } a +N (b +N c) b +N (a +N c)
pr {a} {b} {c} rewrite Semiring.+Associative Semiring a b c | Semiring.commutative Semiring a b | equalityCommutative (Semiring.+Associative Semiring b a c) = refl
lemm : a +N (n +N subtractionNResult.result (-N (inl n<b+c))) n +N n
lemm = identityOfIndiscernablesLeft _≡_ (applyEquality (n +N_) pr1) (pr {n} {a})
lemm2 : a +N (b +N c) n +N n
lemm2 = identityOfIndiscernablesLeft _≡_ lemm (applyEquality (a +N_) (addMinus' (inl n<b+c)))
lemm3 : n +N n <N (n +N subtractionNResult.result (-N (inl n<a+b))) +N c
lemm3 = identityOfIndiscernablesRight _<N_ (additionPreservesInequalityOnLeft n pr2) (Semiring.+Associative Semiring n _ c)
lemm4 : n +N n <N (a +N b) +N c
lemm4 = identityOfIndiscernablesRight _<N_ lemm3 (applyEquality (_+N c) (addMinus' (inl n<a+b)))
lemm5 : n +N n <N a +N (b +N c)
lemm5 = identityOfIndiscernablesRight _<N_ lemm4 (equalityCommutative (Semiring.+Associative Semiring a b c))
lemm6 : a +N (b +N c) <N a +N (b +N c)
lemm6 = identityOfIndiscernablesLeft _<N_ lemm5 (equalityCommutative lemm2)
help9 : {a n : } (a +N 0 n) (a <N n) False
help9 {a} {n} n=a+0 a<n rewrite Semiring.commutative Semiring a 0 | n=a+0 = TotalOrder.irreflexive TotalOrder a<n
help8 : {a n : } (n <N a +N 0) (a <N n) False
help8 {a} {n} n<a+0 a<n rewrite Semiring.commutative Semiring a 0 = TotalOrder.irreflexive TotalOrder (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 Semiring.commutative Semiring a 0 = canSubtractFromEqualityLeft {n} lem'
where
lem : n +N a (n +N subtractionNResult.result (-N (inl n<a+b))) +N c
lem rewrite addMinus' (inl n<a+b) | equalityCommutative (Semiring.+Associative Semiring a b c) | b+c=n = Semiring.commutative Semiring n a
lem' : n +N a n +N (subtractionNResult.result (-N (inl n<a+b)) +N c)
lem' = identityOfIndiscernablesRight _≡_ lem (equalityCommutative (Semiring.+Associative Semiring n _ c))
help5 : {a b c n : } (n<b+c : n <N b +N c) (n<a+b : n <N a +N b) a +N subtractionNResult.result (-N (inl n<b+c)) subtractionNResult.result (-N (inl n<a+b)) +N c
help5 {a} {b} {c} {n} n<b+c n<a+b = canSubtractFromEqualityLeft {n} lemma''
where
lemma : a +N (n +N subtractionNResult.result (-N (inl n<b+c))) (n +N subtractionNResult.result (-N (inl n<a+b))) +N c
lemma rewrite addMinus' (inl n<b+c) | addMinus' (inl n<a+b) = Semiring.+Associative Semiring a b c
lemma' : a +N (n +N subtractionNResult.result (-N (inl n<b+c))) n +N (subtractionNResult.result (-N (inl n<a+b)) +N c)
lemma' = identityOfIndiscernablesRight _≡_ lemma (equalityCommutative (Semiring.+Associative Semiring n _ c))
lemma'' : n +N (a +N subtractionNResult.result (-N (inl n<b+c))) n +N (subtractionNResult.result (-N (inl n<a+b)) +N c)
lemma'' = identityOfIndiscernablesLeft _≡_ lemma' (pr {a} {n} {subtractionNResult.result (-N (inl n<b+c))})
where
pr : {a b c : } a +N (b +N c) b +N (a +N c)
pr {a} {b} {c} rewrite Semiring.+Associative Semiring a b c | Semiring.commutative Semiring a b | equalityCommutative (Semiring.+Associative Semiring b a c) = refl
help4 : {a b c n : } (n<a+'b+c : n <N a +N (b +N c)) (n<a+b : n <N a +N b) (subtractionNResult.result (-N (inl n<a+'b+c)) subtractionNResult.result (-N (inl n<a+b)) +N c)
help4 {a} {b} {c} {n} n<a+'b+c n<a+b = canSubtractFromEqualityLeft lemma'
where
lemma : (n +N subtractionNResult.result (-N (inl n<a+'b+c))) (n +N subtractionNResult.result (-N (inl n<a+b))) +N c
lemma rewrite addMinus' (inl n<a+'b+c) | addMinus' (inl n<a+b) = Semiring.+Associative Semiring a b c
lemma' : n +N subtractionNResult.result (-N (inl n<a+'b+c)) n +N (subtractionNResult.result (-N (inl n<a+b)) +N c)
lemma' = identityOfIndiscernablesRight _≡_ lemma (equalityCommutative (Semiring.+Associative Semiring n (subtractionNResult.result (-N (inl n<a+b))) c))
help3 : {a b c n : } (a <N n) (b <N n) (c <N n) (a +N b <N n) (pr : n <N b +N c) a +N subtractionNResult.result (-N (inl pr)) n False
help3 {a} {b} {c} {n} a<n b<n c<n a+b<n n<b+c pr = TotalOrder.irreflexive TotalOrder (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 Semiring.+Associative Semiring a b c | Semiring.+Associative Semiring b a c = applyEquality (_+N c) (Semiring.commutative Semiring a b)
inter2 : n +N n a +N (b +N c)
inter2 = equalityCommutative (identityOfIndiscernablesLeft _≡_ inter (applyEquality (a +N_) (addMinus' (inl n<b+c))))
inter3 : n +N n <N n +N c
inter3 rewrite inter2 | Semiring.+Associative Semiring a b c = additionPreservesInequality c a+b<n
inter4 : (n +N n <N n +N c) n <N c
inter4 pr rewrite Semiring.commutative Semiring n c = subtractionPreservesInequality n pr
help2 : {a b c n : } (sn<b+c : succ n <N b +N c) (sn<a+b+c : succ n <N (a +N b) +N c) a +N subtractionNResult.result (-N (inl sn<b+c)) subtractionNResult.result (-N (inl sn<a+b+c))
help2 {a} {b} {c} {n} sn<b+c sn<a+b+c = res inter
where
inter : a +N (subtractionNResult.result (-N (inl sn<b+c)) +N succ n) subtractionNResult.result (-N (inl sn<a+b+c)) +N succ n
inter rewrite addMinus (inl sn<b+c) | addMinus (inl sn<a+b+c) = Semiring.+Associative Semiring a b c
res : (a +N (subtractionNResult.result (-N (inl sn<b+c)) +N succ n) subtractionNResult.result (-N (inl sn<a+b+c)) +N succ n) a +N subtractionNResult.result (-N (inl sn<b+c)) subtractionNResult.result (-N (inl sn<a+b+c))
res pr = canSubtractFromEqualityRight {_} {succ n} (identityOfIndiscernablesLeft _≡_ pr (Semiring.+Associative Semiring a (subtractionNResult.result (-N (inl sn<b+c))) (succ n)))
help1 : {a b c n : } (sn<b+c : succ n <N b +N c) (pr1 : succ n <N a +N subtractionNResult.result (-N (inl sn<b+c))) (a +N b <N succ n) (a <N succ n) (b <N succ n) (c <N succ n) False
help1 {a} {b} {c} {n} sn<b+c pr1 a+b<sn a<sn b<sn c<sn with -N (inl sn<b+c)
help1 {a} {b} {c} {n} sn<b+c pr1 a+b<sn a<sn b<sn c<sn | record { result = b+c-sn ; pr = Prb+c-sn } = ans
where
b+c-nNonzero : b+c-sn 0 False
b+c-nNonzero pr rewrite (equalityCommutative Prb+c-sn) | pr | Semiring.commutative Semiring n 0 = TotalOrder.irreflexive TotalOrder sn<b+c
2sn<a+b+c' : succ n +N succ n <N succ n +N (a +N b+c-sn)
2sn<a+b+c' = additionPreservesInequalityOnLeft (succ n) pr1
2sn<a+b+c'' : succ n +N succ n <N a +N (succ n +N b+c-sn)
2sn<a+b+c'' rewrite Semiring.+Associative Semiring a (succ n) b+c-sn | Semiring.commutative Semiring a (succ n) | equalityCommutative (Semiring.+Associative Semiring (succ n) a b+c-sn) = 2sn<a+b+c'
eep : succ n +N succ n <N a +N (b +N c)
eep rewrite equalityCommutative Prb+c-sn = 2sn<a+b+c''
eep2 : a +N (b +N c) <N succ n +N c
eep2 rewrite Semiring.+Associative Semiring a b c = additionPreservesInequality c a+b<sn
eep2' : a +N (b +N c) <N succ n +N succ n
eep2' = orderIsTransitive eep2 (additionPreservesInequalityOnLeft (succ n) c<sn)
ans : False
ans = TotalOrder.irreflexive TotalOrder (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 _ _ (Semiring.+Associative Semiring a b c)
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inl b+c<sn) | inl (inr sn<a+'b+c) = exFalso (false {succ n} a+b+c<sn sn<a+'b+c)
where
false : {x : } (a +N b) +N c <N succ n succ n <N a +N (b +N c) False
false pr1 pr2 rewrite equalityCommutative (Semiring.+Associative Semiring a b c) = TotalOrder.irreflexive TotalOrder (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 equalityCommutative (Semiring.+Associative Semiring a b c) | p2 = TotalOrder.irreflexive TotalOrder p1
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inr sn<b+c) with 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 | Semiring.+Associative Semiring a (succ n) result | Semiring.commutative Semiring a (succ n) | equalityCommutative (Semiring.+Associative Semiring (succ n) a result) = cannotAddAndEnlarge' pr1
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inr sn<b+c) | inl (inl _) | inl (inr x) with -N (inl sn<b+c)
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inr sn<b+c) | inl (inl a+'b+c<sn) | inl (inr x) | record { result = result ; pr = pr } = exFalso false
where
lemma : a +N (succ n +N result) a +N (b +N c)
lemma' : a +N (succ n +N result) <N succ n
lemma'' : succ n +N (a +N result) <N succ n
lemma'' = identityOfIndiscernablesLeft _<N_ lemma' (transitivity (Semiring.+Associative Semiring a (succ n) result) (transitivity (applyEquality (λ t t +N result) (Semiring.commutative Semiring a (succ n))) (equalityCommutative (Semiring.+Associative Semiring (succ n) a result))))
lemma = applyEquality (λ i a +N i) pr
lemma' rewrite lemma = a+'b+c<sn
false : False
false = cannotAddAndEnlarge' lemma''
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inr sn<b+c) | inl (inl _) | inr x with -N (inl sn<b+c)
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inr sn<b+c) | inl (inl a+'b+c<sn) | inr x | record { result = result ; pr = pr } = exFalso false
where
lemma : a +N (succ n +N result) <N succ n
lemma rewrite pr = a+'b+c<sn
lemma' : succ n +N (a +N result) <N succ n
lemma' = identityOfIndiscernablesLeft _<N_ lemma (transitivity (Semiring.+Associative Semiring a (succ n) result) (transitivity (applyEquality (λ t t +N result) (Semiring.commutative Semiring a (succ n))) (equalityCommutative (Semiring.+Associative Semiring (succ n) a result))))
false : False
false = cannotAddAndEnlarge' lemma'
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inr sn<b+c) | inl (inr x) = equalityZn _ _ (exFalso (false {succ n} a+b+c<sn x))
where
false : {x : } (a +N b) +N c <N succ n succ n <N a +N (b +N c) False
false pr1 pr2 rewrite equalityCommutative (Semiring.+Associative Semiring a b c) = TotalOrder.irreflexive TotalOrder (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 | equalityCommutative (Semiring.+Associative Semiring a b c) = TotalOrder.irreflexive TotalOrder pr1
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inr sn=b+c = exFalso (false a+b+c<sn sn=b+c)
where
false : (a +N b) +N c <N succ n b +N c succ n False
false pr1 pr2 rewrite equalityCommutative (Semiring.+Associative Semiring a b c) | pr2 | Semiring.commutative Semiring a (succ n) = cannotAddAndEnlarge' a+b+c<sn
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) with 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 equalityCommutative (Semiring.+Associative Semiring a b c) = TotalOrder.irreflexive TotalOrder (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 equalityCommutative (Semiring.+Associative Semiring a b c) = Semiring.commutative Semiring (succ n) _
ans : subtractionNResult.result (-N (inl sn<a+'b+c)) subtractionNResult.result (-N (inl sn<a+b+c))
ans = equivalentSubtraction _ _ _ _ sn<a+'b+c sn<a+b+c lemma
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inl (inl b+c<sn) | inr x = exFalso (false sn<a+b+c x)
where
false : succ n <N (a +N b) +N c (a +N (b +N c)) succ n False
false pr1 pr2 rewrite equalityCommutative (Semiring.+Associative Semiring a b c) | pr2 = TotalOrder.irreflexive TotalOrder sn<a+b+c
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inl (inr sn<b+c) with 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 equalityCommutative (Semiring.+Associative Semiring a b c) = TotalOrder.irreflexive TotalOrder (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 equalityCommutative (Semiring.+Associative Semiring a b c) | pr2 = TotalOrder.irreflexive TotalOrder pr1
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inr sn=b+c with 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 equalityCommutative (Semiring.+Associative Semiring a b c) = TotalOrder.irreflexive TotalOrder (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 Semiring.commutative Semiring a 0 | equalityCommutative (Semiring.+Associative Semiring a b c) | b+c=sn | Semiring.commutative Semiring (succ n) result = equalityCommutative (canSubtractFromEqualityRight pr1)
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inr sn=b+c | inl (inr _) | inl (inr x) = exFalso (false b a+b<sn x)
where
false : (b : ) a +N b <N succ n succ n <N a +N 0 False
false zero pr1 pr2 rewrite Semiring.commutative Semiring a 0 = TotalOrder.irreflexive TotalOrder (orderIsTransitive pr1 pr2)
false (succ b) pr1 pr2 rewrite Semiring.commutative Semiring a 0 = TotalOrder.irreflexive TotalOrder (orderIsTransitive pr2 (orderIsTransitive (le b (Semiring.commutative Semiring (succ b) a)) pr1))
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inr sn=b+c | inl (inr _) | inr x = exFalso (false b a+b<sn x)
where
false : (b : ) a +N b <N succ n a +N 0 succ n False
false zero pr1 pr2 rewrite pr2 = TotalOrder.irreflexive TotalOrder pr1
false (succ b) pr1 pr2 rewrite Semiring.commutative Semiring a 0 | pr2 = cannotAddAndEnlarge' pr1
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inr sn=b+c | inr x = exFalso (false sn<a+b+c x)
where
false : succ n <N (a +N b) +N c a +N (b +N c) succ n False
false pr1 pr2 rewrite equalityCommutative (Semiring.+Associative Semiring a b c) | pr2 = TotalOrder.irreflexive TotalOrder pr1
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inr sn=a+b+c with 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 equalityCommutative (Semiring.+Associative Semiring a b c) | pr1 = TotalOrder.irreflexive TotalOrder pr2
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inr sn=a+b+c | inl (inl b+c<sn) | inl (inr sn<a+'b+c) = exFalso (false sn=a+b+c sn<a+'b+c)
where
false : (a +N b) +N c succ n succ n <N a +N (b +N c) False
false pr1 pr2 rewrite equalityCommutative (Semiring.+Associative Semiring a b c) | pr1 = TotalOrder.irreflexive TotalOrder pr2
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inr sn=a+b+c | inl (inl b+c<sn) | inr _ = equalityZn _ _ refl
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inr sn=a+b+c | inl (inr sn<b+c) = exFalso (false a sn=a+b+c sn<b+c)
where
false : (a : ) (a +N b) +N c succ n succ n <N b +N c False
false zero pr1 pr2 rewrite equalityCommutative (Semiring.+Associative Semiring a b c) | equalityCommutative pr1 = TotalOrder.irreflexive TotalOrder pr2
false (succ a) pr1 pr2 rewrite equalityCommutative (Semiring.+Associative Semiring a b c) | equalityCommutative pr1 = TotalOrder.irreflexive TotalOrder (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 equalityCommutative (Semiring.+Associative Semiring a b c) | pr2 = refl
a=0 (succ a) pr1 pr2 rewrite equalityCommutative (Semiring.+Associative Semiring a b c) | pr2 = canSubtractFromEqualityRight pr1
ans : a +N 0 0
ans rewrite Semiring.commutative Semiring a 0 = a=0 a sn=a+b+c b+c=sn
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inr sn=a+b+c | inr b+c=sn | inl (inr sn<a+0) = exFalso (false sn<a+0 sn=a+b+c b+c=sn)
where
false : succ n <N a +N 0 (a +N b) +N c succ n b +N c succ n False
false pr1 pr2 pr3 rewrite equalityCommutative (Semiring.+Associative Semiring a b c) | pr3 | Semiring.commutative Semiring a 0 = zeroNeverGreater {succ n} (identityOfIndiscernablesRight _<N_ pr1 (a=0 a pr2))
where
a=0 : (a : ) (a +N succ n succ n) a 0
a=0 zero pr = refl
a=0 (succ a) pr = canSubtractFromEqualityRight pr
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inr sn=a+b+c | inr b+c=sn | inr sn=a+0 = exFalso (false sn=a+b+c b+c=sn sn=a+0)
where
false : (a +N b) +N c succ n b +N c succ n a +N 0 succ n False
false pr1 pr2 pr3 rewrite equalityCommutative (Semiring.+Associative Semiring a b c) | pr2 | equalityCommutative pr3 | Semiring.commutative Semiring a 0 = naughtE (identityOfIndiscernablesLeft _≡_ pr3 (a=0 a pr1))
where
a=0 : (a : ) (a +N a a) a 0
a=0 zero pr = refl
a=0 (succ a) pr = exFalso (naughtE {a} (equalityCommutative (canSubtractFromEqualityRight pr)))
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) with 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 Semiring.+Associative Semiring 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 Semiring.+Associative Semiring a b c | equalityCommutative pr2 = cannotAddAndEnlarge' pr1
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inl (inr sn<b+c) with 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 (TotalOrder.irreflexive TotalOrder (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 | Semiring.commutative Semiring b c = cannotAddAndEnlarge' b+c<sn
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inr c=sn | inl (inr sn<b+c) with 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 (TotalOrder.irreflexive TotalOrder 0<a)
subLess {succ a} {b} _ a<b with -N (inl a<b)
... | record { result = b-a ; pr = pr } = record { x = a ; proof = pr }
inverseZn : {n : } {pr : 0 <N n} (a : n n pr) Sg (n n pr) (λ i i +n a record { x = 0 ; xLess = pr } )
inverseZn {zero} {()}
inverseZn {succ n} {0<n} record { x = zero ; xLess = zeroLess } = record { x = zero ; xLess = zeroLess } , plusZnIdentityLeft _
inverseZn {succ n} {0<n} record { x = (succ a) ; xLess = aLess } = ans , pr
where
ans = record { x = subtractionNResult.result (-N (inl (canRemoveSuccFrom<N aLess))) ; xLess = subLess (succIsPositive a) aLess }
pr : ans +n record { x = (succ a) ; xLess = aLess } record { x = 0 ; xLess = 0<n }
pr with orderIsTotal (subtractionNResult.result (-N (inl (canRemoveSuccFrom<N aLess))) +N (succ a)) (succ n)
... | inl (inl x) = exFalso f
where
h : subtractionNResult.result (-N (inl (canRemoveSuccFrom<N aLess))) +N succ a succ n
h with -N (inl (canRemoveSuccFrom<N aLess))
h | record { result = result ; pr = pr } rewrite equalityCommutative pr = Semiring.commutative Semiring result (succ a)
f : False
f = TotalOrder.irreflexive TotalOrder (identityOfIndiscernablesLeft _<N_ x h)
... | inl (inr x) = exFalso f
where
h : subtractionNResult.result (-N (inl (canRemoveSuccFrom<N aLess))) +N succ a succ n
h = addMinus {succ a} {succ n} (inl aLess)
f : False
f rewrite h = TotalOrder.irreflexive TotalOrder x
... | (inr n-a+sa=sn) = equalityZn _ _ refl
nGroup : (n : ) (pr : 0 <N n) Group (reflSetoid (n n pr)) _+n_
nGroup n pr = record { 0G = record { x = 0 ; xLess = pr } ; inverse = λ a underlying (inverseZn a) ; +Associative = λ {a} {b} {c} plusZnAssociative a b c ; identRight = λ {a} plusZnIdentityRight a ; identLeft = λ {a} plusZnIdentityLeft a ; invLeft = λ {a} helpInvLeft a ; invRight = λ {a} helpInvRight a ; +WellDefined = reflGroupWellDefined }
where
helpInvLeft : (a : n n pr) underlying (inverseZn a) +n a record { x = 0 ; xLess = pr }
helpInvLeft a with inverseZn a
... | -a , pr-a = pr-a
helpInvRight : (a : n n pr) a +n underlying (inverseZn a) record { x = 0 ; xLess = pr }
helpInvRight a rewrite plusZnCommutative a (underlying (inverseZn a)) = helpInvLeft a
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

View File

@@ -5,6 +5,7 @@ open import Numbers.Naturals.Naturals
open import Numbers.Integers.Integers
open import Groups.Groups
open import Groups.Definition
open import Groups.Lemmas
open import Rings.Definition
open import Rings.Orders.Total.Definition
open import Rings.Orders.Partial.Definition
@@ -15,7 +16,7 @@ open import Setoids.Orders
open import Functions
open import Sets.EquivalenceRelations
module Numbers.Rationals where
module Numbers.Rationals.Definition where
open import Fields.FieldOfFractions.Setoid IntDom
open import Fields.FieldOfFractions.Addition IntDom

View File

@@ -16,11 +16,11 @@ open import Setoids.Setoids
open import Setoids.Orders
open import Functions
open import Sets.EquivalenceRelations
open import Numbers.Rationals
open import Numbers.Rationals.Definition
open import Semirings.Definition
open import Orders
module Numbers.RationalsLemmas where
module Numbers.Rationals.Lemmas where
open PartiallyOrderedRing POrderedRing
open import Rings.Orders.Total.Lemmas OrderedRing

View File

@@ -2,6 +2,7 @@
open import LogicalFormulae
open import Groups.Groups
open import Groups.Homomorphisms.Definition
open import Groups.Definition
open import Numbers.Naturals.Naturals
open import Setoids.Orders

View File

@@ -5,7 +5,8 @@ open import Groups.Groups
open import Functions
open import Numbers.Naturals.Naturals
open import Numbers.Integers.Integers
open import Numbers.Modulo.IntegersModN
open import Numbers.Modulo.Group
open import Numbers.Modulo.Definition
open import Rings.Examples.Proofs
open import Numbers.Primes.PrimeNumbers

View File

@@ -9,7 +9,8 @@ open import Rings.Definition
open import Numbers.Naturals.Naturals
open import Numbers.Integers.Integers
open import Numbers.Primes.PrimeNumbers
open import Numbers.Modulo.IntegersModN
open import Numbers.Modulo.Definition
open import Numbers.Modulo.Group
module Rings.Examples.Proofs where
nToZn' : (n : ) (pr : 0 <N n) (x : ) n n pr

View File

@@ -2,6 +2,7 @@
open import LogicalFormulae
open import Groups.Groups
open import Groups.Lemmas
open import Groups.Definition
open import Numbers.Naturals.Naturals
open import Orders

View File

@@ -50,7 +50,7 @@ abstract
lemm3 a b pr1 pr2 with transferToRight' additiveGroup (Equivalence.symmetric eq pr1)
... | a=-b with Equivalence.transitive eq pr2 a=-b
... | 0=-b with inverseWellDefined additiveGroup 0=-b
... | -0=--b = Equivalence.transitive eq (Equivalence.symmetric eq (invIdentity additiveGroup)) (Equivalence.transitive eq -0=--b (invTwice additiveGroup b))
... | -0=--b = Equivalence.transitive eq (Equivalence.symmetric eq (invIdent additiveGroup)) (Equivalence.transitive eq -0=--b (invTwice additiveGroup b))
charNot2ImpliesNontrivial : ((1R + 1R) 0R False) (0R 1R) False
charNot2ImpliesNontrivial charNot2 0=1 = charNot2 (Equivalence.transitive eq (+WellDefined (Equivalence.symmetric eq 0=1) (Equivalence.symmetric eq 0=1)) identRight)

View File

@@ -90,10 +90,10 @@ abstract
triangleInequality a b | inl (inr a+b<0) | inl (inr a<0) with totality 0G b
triangleInequality a b | inl (inr a+b<0) | inl (inr a<0) | inl (inl 0<b) = inl (<WellDefined (Equivalence.symmetric eq (invContravariant additiveGroup)) groupIsAbelian (orderRespectsAddition (SetoidPartialOrder.<Transitive pOrder (lemm2' _ 0<b) 0<b) (inverse a)))
triangleInequality a b | inl (inr a+b<0) | inl (inr a<0) | inl (inr b<0) = inr (Equivalence.transitive eq (invContravariant additiveGroup) groupIsAbelian)
triangleInequality a b | inl (inr a+b<0) | inl (inr a<0) | inr 0=b = inr (Equivalence.transitive eq (invContravariant additiveGroup) (Equivalence.transitive eq (Equivalence.transitive eq (Equivalence.transitive eq (+WellDefined (Equivalence.transitive eq (inverseWellDefined additiveGroup (Equivalence.symmetric eq 0=b)) (invIdentity additiveGroup)) (Equivalence.reflexive eq)) identLeft) (Equivalence.symmetric eq identRight)) (+WellDefined (Equivalence.reflexive eq) 0=b)))
triangleInequality a b | inl (inr a+b<0) | inl (inr a<0) | inr 0=b = inr (Equivalence.transitive eq (invContravariant additiveGroup) (Equivalence.transitive eq (Equivalence.transitive eq (Equivalence.transitive eq (+WellDefined (Equivalence.transitive eq (inverseWellDefined additiveGroup (Equivalence.symmetric eq 0=b)) (invIdent additiveGroup)) (Equivalence.reflexive eq)) identLeft) (Equivalence.symmetric eq identRight)) (+WellDefined (Equivalence.reflexive eq) 0=b)))
triangleInequality a b | inl (inr a+b<0) | inr 0=a with totality 0G b
triangleInequality a b | inl (inr a+b<0) | inr 0=a | inl (inl 0<b) = exFalso (irreflexive {0G} (SetoidPartialOrder.<Transitive pOrder 0<b (<WellDefined (Equivalence.transitive eq (+WellDefined (Equivalence.symmetric eq 0=a) (Equivalence.reflexive eq)) identLeft) (Equivalence.reflexive eq) a+b<0)))
triangleInequality a b | inl (inr a+b<0) | inr 0=a | inl (inr b<0) = inr (Equivalence.transitive eq (invContravariant additiveGroup) (Equivalence.transitive eq groupIsAbelian (+WellDefined (Equivalence.transitive eq (Equivalence.transitive eq (Equivalence.symmetric eq (inverseWellDefined additiveGroup 0=a)) (invIdentity additiveGroup)) 0=a) (Equivalence.reflexive eq))))
triangleInequality a b | inl (inr a+b<0) | inr 0=a | inl (inr b<0) = inr (Equivalence.transitive eq (invContravariant additiveGroup) (Equivalence.transitive eq groupIsAbelian (+WellDefined (Equivalence.transitive eq (Equivalence.transitive eq (Equivalence.symmetric eq (inverseWellDefined additiveGroup 0=a)) (invIdent additiveGroup)) 0=a) (Equivalence.reflexive eq))))
triangleInequality a b | inl (inr a+b<0) | inr 0=a | inr 0=b = exFalso (irreflexive {0G} (<WellDefined (Equivalence.transitive eq (+WellDefined (Equivalence.symmetric eq 0=a) (Equivalence.symmetric eq 0=b)) identLeft) (Equivalence.reflexive eq) a+b<0))
triangleInequality a b | inr 0=a+b with totality 0G a
triangleInequality a b | inr 0=a+b | inl (inl 0<a) with totality 0G b
@@ -169,7 +169,7 @@ abstract
f : inverse 0G inverse (y + inverse x)
f = inverseWellDefined additiveGroup 0=y-x
g : 0G (inverse y) + x
g = Equivalence.transitive eq (symmetric (invIdentity additiveGroup)) (Equivalence.transitive eq f (Equivalence.transitive eq (Equivalence.transitive eq (invContravariant additiveGroup) groupIsAbelian) (+WellDefined reflexive (invInv additiveGroup))))
g = Equivalence.transitive eq (symmetric (invIdent additiveGroup)) (Equivalence.transitive eq f (Equivalence.transitive eq (Equivalence.transitive eq (invContravariant additiveGroup) groupIsAbelian) (+WellDefined reflexive (invInv additiveGroup))))
x=y : x y
x=y = transferToRight additiveGroup (symmetric (Equivalence.transitive eq g groupIsAbelian))
q'' : (0R + x) < ((y + Group.inverse additiveGroup x) + x)