Z is a Euclidean domain (#86)

This commit is contained in:
Patrick Stevens
2019-12-07 13:00:18 +00:00
committed by GitHub
parent cfd9787bb8
commit e192f0e1f1
38 changed files with 1018 additions and 486 deletions

View File

@@ -2,11 +2,14 @@
open import LogicalFormulae
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Numbers.Naturals.Naturals
open import Numbers.Naturals.Semiring
open import Numbers.Naturals.Order
open import Vectors
open import Semirings.Definition
open import Categories.Definition
open import Orders
open import Categories.Functor.Definition
open import Categories.Examples
module Categories.Category where
@@ -20,7 +23,7 @@ NatPreorder : Category {lzero} {lzero}
NatPreorder = record { objects = ; arrows = λ m n m ≤N n ; id = λ x inr refl ; _∘_ = λ f g leqTransitive g f ; rightId = λ x<y leqUnique (leqTransitive x<y (inr refl)) x<y ; leftId = λ x<y leqUnique (leqTransitive (inr refl) x<y) x<y ; associative = λ z<=w y<=z x<=y leqUnique (leqTransitive (leqTransitive x<=y y<=z) z<=w) (leqTransitive x<=y (leqTransitive y<=z z<=w)) }
where
leqTransitive : {a b c : } (a ≤N b) (b ≤N c) (a ≤N c)
leqTransitive (inl a<b) (inl b<c) = inl (orderIsTransitive a<b b<c)
leqTransitive (inl a<b) (inl b<c) = inl (TotalOrder.<Transitive TotalOrder a<b b<c)
leqTransitive (inl a<b) (inr b=c) rewrite b=c = inl a<b
leqTransitive (inr a=b) (inl b<c) rewrite a=b = inl b<c
leqTransitive (inr a=b) (inr b=c) rewrite a=b | b=c = inr refl

View File

@@ -9,6 +9,7 @@ open import Numbers.BinaryNaturals.Order
open import Numbers.BinaryNaturals.Subtraction
open import Numbers.Integers.Integers
open import Numbers.Integers.RingStructure.EuclideanDomain
open import Lists.Lists
@@ -33,6 +34,7 @@ open import Groups.Cyclic.Definition
open import Groups.Cyclic.DefinitionLemmas
open import Groups.Polynomials.Examples
open import Groups.Cosets
open import Groups.Subgroups.Normal.Examples
open import Fields.Fields
open import Fields.Orders.Partial.Definition
@@ -65,9 +67,13 @@ open import Rings.Ideals.Prime.Definition
open import Rings.Ideals.Prime.Lemmas
open import Rings.Ideals.Principal.Definition
open import Rings.IntegralDomains.Examples
open import Rings.PrincipalIdealDomain
open import Rings.PrincipalIdealDomains.Definition
open import Rings.Subrings.Definition
open import Rings.Ideals.Maximal.Lemmas
open import Rings.Primes.Lemmas
open import Rings.Irreducibles.Definition
open import Rings.Divisible.Definition
open import Rings.Associates.Lemmas
open import Setoids.Setoids
open import Setoids.DirectSum

7
Everything/Unsafe.agda Normal file
View File

@@ -0,0 +1,7 @@
{-# OPTIONS --warning=error --allow-unsolved-metas #-}
-- This file contains everything that cannot be compiled in --safe mode.
open import Lists.SortList
module Everything.Unsafe where

View File

@@ -11,12 +11,12 @@ open import Functions
open import Setoids.Setoids
open import Sets.EquivalenceRelations
open import Rings.IntegralDomains.Definition
open import Rings.IntegralDomains.Lemmas
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
module Fields.FieldOfFractions.Setoid {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A A A} {_*_ : A A A} {R : Ring S _+_ _*_} (I : IntegralDomain R) where
fieldOfFractionsSet : Set (a b)
fieldOfFractionsSet = (A && (Sg A (λ a (Setoid.__ S a (Ring.0R R) False))))
@@ -38,6 +38,6 @@ Equivalence.transitive (Setoid.eq fieldOfFractionsSetoid) {a ,, (b , b!=0)} {c ,
p3 : (a * f) * d (b * e) * d
p3 = transitive p2 (transitive (*WellDefined reflexive *Commutative) *Associative)
p4 : ((d 0R) False) ((a * f) (b * e))
p4 = cancelIntDom R I (transitive *Commutative (transitive p3 *Commutative))
p4 = cancelIntDom I (transitive *Commutative (transitive p3 *Commutative))
p5 : (a * f) (b * e)
p5 = p4 d!=0

View File

@@ -13,6 +13,7 @@ open import Setoids.Orders
open import Functions
open import LogicalFormulae
open import Numbers.Naturals.Naturals
open import Rings.IntegralDomains.Definition
module Fields.Lemmas {m n : _} {A : Set m} {S : Setoid {m} {n} A} {_+_ : A A A} {_*_ : A A A} {R : Ring S _+_ _*_} (F : Field R) where
@@ -32,3 +33,10 @@ abstract
halfHalves : {x : A} (1/2 : A) (pr : 1/2 + 1/2 1R) (x + x) * 1/2 x
halfHalves {x} 1/2 pr = Equivalence.transitive eq (Equivalence.transitive eq (Equivalence.transitive eq *Commutative (Equivalence.transitive eq (Equivalence.transitive eq *DistributesOver+ (Equivalence.transitive eq (+WellDefined *Commutative *Commutative) (Equivalence.symmetric eq *DistributesOver+))) *Commutative)) (*WellDefined pr (Equivalence.reflexive eq))) identIsIdent
fieldIsIntDom : (Setoid.__ S (Ring.1R R) (Ring.0R R) False) IntegralDomain R
IntegralDomain.intDom (fieldIsIntDom 1!=0) {a} {b} ab=0 a!=0 with Field.allInvertible F a a!=0
IntegralDomain.intDom (fieldIsIntDom _) {a} {b} ab=0 a!=0 | 1/a , prA = transitive (symmetric identIsIdent) (transitive (*WellDefined (symmetric prA) reflexive) (transitive (symmetric *Associative) (transitive (*WellDefined reflexive ab=0) (Ring.timesZero R))))
where
open Equivalence eq
IntegralDomain.nontrivial (fieldIsIntDom 1!=0) = 1!=0

View File

@@ -5,12 +5,13 @@ open import Setoids.Setoids
open import Sets.EquivalenceRelations
open import Functions
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Numbers.Naturals.Naturals
open import Numbers.Naturals.Definition
open import Numbers.Integers.Integers
open import Numbers.Integers.Addition
open import Sets.FinSet
open import Groups.Homomorphisms.Definition
open import Groups.Groups
open import Groups.Subgroups.Definition
open import Groups.Lemmas
open import Groups.Subgroups.Definition
open import Groups.Abelian.Definition
@@ -18,21 +19,24 @@ open import Groups.Definition
open import Groups.Cyclic.Definition
open import Groups.Cyclic.DefinitionLemmas
open import Semirings.Definition
open import Rings.Definition
module Groups.Cyclic.Lemmas where
module Groups.Cyclic.Lemmas {m n : _} {A : Set m} {S : Setoid {m} {n} A} {_·_ : A A A} (G : Group S _·_) where
elementPowerHom : {m n : _} {A : Set m} {S : Setoid {m} {n} A} {_·_ : A A A} (G : Group S _·_) (x : A) GroupHom Group G (λ i elementPower G x i)
GroupHom.groupHom (elementPowerHom {S = S} G x) {a} {b} = symmetric (elementPowerCollapse G x a b)
elementPowerHom : (x : A) GroupHom Group G (λ i elementPower G x i)
GroupHom.groupHom (elementPowerHom x) {a} {b} = symmetric (elementPowerCollapse G x a b)
where
open Equivalence (Setoid.eq S)
GroupHom.wellDefined (elementPowerHom {S = S} G x) {.y} {y} refl = reflexive
GroupHom.wellDefined (elementPowerHom x) {.y} {y} refl = reflexive
where
open Equivalence (Setoid.eq S)
subgroupOfCyclicIsCyclic : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+A_ : A A A} {_+B_ : B B B} {G : Group S _+A_} {H : Group T _+B_} {f : B A} {fHom : GroupHom H G f} Subgroup G H fHom CyclicGroup G CyclicGroup H
CyclicGroup.generator (subgroupOfCyclicIsCyclic {f = f} subgrp record { generator = generator ; cyclic = cyclic }) = {!!}
CyclicGroup.cyclic (subgroupOfCyclicIsCyclic {f = f} subgrp record { generator = generator ; cyclic = cyclic }) {a} with cyclic {f a}
CyclicGroup.cyclic (subgroupOfCyclicIsCyclic {f = f} subgrp record { generator = generator ; cyclic = cyclic }) {a} | N , generator^n=fa = N , {!!}
subgroupOfCyclicIsCyclic : {c : _} {pred : A Set c} (sub : Subgroup G pred) CyclicGroup G CyclicGroup (subgroupIsGroup G sub)
CyclicGroup.generator (subgroupOfCyclicIsCyclic {pred = pred} sub record { generator = g ; cyclic = cyclic }) = {!!}
where
leastPowerInGroup : (bound : )
leastPowerInGroup bound = {!!}
CyclicGroup.cyclic (subgroupOfCyclicIsCyclic sub cyc) = {!!}
-- Prefer to prove that subgroup of cyclic is cyclic, then deduce this like with abelian groups
{-
@@ -43,6 +47,11 @@ CyclicGroup.cyclic (cyclicIsGroupProperty {H = H} iso G) {a} | a' , b with Cycli
... | pow , prPow = pow , {!!}
-}
-- Proof of abelianness of cyclic groups: a cyclic group is the image of elementPowerHom into Z, so is isomorphic to a subgroup of Z. All subgroups of an abelian group are abelian.
cyclicGroupIsAbelian : {m n : _} {A : Set m} {S : Setoid {m} {n} A} {_+_ : A A A} {G : Group S _+_} (cyclic : CyclicGroup G) AbelianGroup G
cyclicGroupIsAbelian cyclic = {!!}
cyclicGroupIsAbelian : (cyclic : CyclicGroup G) AbelianGroup G
AbelianGroup.commutative (cyclicGroupIsAbelian record { generator = generator ; cyclic = cyclic }) {a} {b} with cyclic {a}
... | bl with cyclic {b}
AbelianGroup.commutative (cyclicGroupIsAbelian record { generator = generator ; cyclic = cyclic }) {a} {b} | nA , prA | nB , prB = transitive (+WellDefined (symmetric prA) (symmetric prB)) (transitive (symmetric (GroupHom.groupHom (elementPowerHom generator) {nA} {nB})) (transitive (transitive (elementPowerWellDefinedZ' G (nA +Z nB) (nB +Z nA) (Ring.groupIsAbelian Ring {nA} {nB}) {generator}) (GroupHom.groupHom (elementPowerHom generator) {nB} {nA})) (+WellDefined prB prA)))
where
open Setoid S
open Equivalence eq
open Group G

View File

@@ -10,7 +10,9 @@ open import Setoids.Setoids
open import LogicalFormulae
open import Sets.FinSet
open import Functions
open import Numbers.Naturals.Naturals
open import Numbers.Naturals.Definition
open import Numbers.Naturals.Order
open import Numbers.Naturals.Semiring
open import Numbers.Naturals.WithK
open import Numbers.Modulo.Definition
open import Numbers.Modulo.Addition
@@ -54,7 +56,7 @@ CyclicGroup.cyclic Cyclic {negSucc x} = (negSucc x , ans)
elementPowZn : (n : ) {pr : 0 <N succ (succ n)} (power : ) (powerLess : power <N succ (succ n)) {p : 1 <N succ (succ n)} elementPower (nGroup (succ (succ n)) pr) (record { x = 1 ; xLess = p }) (nonneg power) record { x = power ; xLess = powerLess }
elementPowZn n zero powerLess = equalityZn _ _ refl
elementPowZn n {pr} (succ power) powerLess {p} with orderIsTotal (n.x (elementPower (nGroup (succ (succ n)) pr) (record { x = 1 ; xLess = p }) (nonneg power))) (succ n)
elementPowZn n {pr} (succ power) powerLess {p} with TotalOrder.totality TotalOrder (n.x (elementPower (nGroup (succ (succ n)) pr) (record { x = 1 ; xLess = p }) (nonneg power))) (succ n)
elementPowZn n {pr} (succ power) powerLess {p} | inl (inl x) = equalityZn _ _ (applyEquality succ v)
where
t : elementPower (nGroup (succ (succ n)) pr) (record { x = 1 ; xLess = succPreservesInequality (succIsPositive n) }) (nonneg power) record { x = power ; xLess = PartialOrder.<Transitive (TotalOrder.order TotalOrder) (a<SuccA power) powerLess }
@@ -102,7 +104,7 @@ intoZn {succ n} {0<n} x x<n | record { quot = quot ; rem = rem ; pr = pr ; remIs
ans rewrite multiplicationNIsCommutative n 0 = refl
intoZn {succ n} {0<n} x x<n | record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = inr () ; quotSmall = quotSmall }
quotientZn : (n : ) (pr : 0 <N n) GroupIso (quotientGroup Group (modNExampleGroupHom n pr)) (nGroup n pr) (mod n pr)
quotientZn : (n : ) (pr : 0 <N n) GroupIso (quotientGroupByHom Group (modNExampleGroupHom n pr)) (nGroup n pr) (mod n pr)
GroupHom.groupHom (GroupIso.groupHom (quotientZn n pr)) = GroupHom.groupHom (modNExampleGroupHom n pr)
GroupHom.wellDefined (GroupIso.groupHom (quotientZn n pr)) {x} {y} x~y = need
where
@@ -135,23 +137,3 @@ GroupHom.groupHom (trivialGroupHom {S = S} G) = symmetric identRight
open Group G
open Equivalence eq
GroupHom.wellDefined (trivialGroupHom {S = S} G) _ = Equivalence.reflexive (Setoid.eq S)
trivialSubgroupIsNormal : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+A_ : A A A} (G : Group S _+A_) NormalSubgroup G trivialGroup (trivialGroupHom G)
SetoidInjection.wellDefined (Subgroup.fInj (NormalSubgroup.subgroup (trivialSubgroupIsNormal {S = S} G))) {x} {.x} refl = Equivalence.reflexive (Setoid.eq S)
SetoidInjection.injective (Subgroup.fInj (NormalSubgroup.subgroup (trivialSubgroupIsNormal G))) {fzero} {fzero} _ = refl
SetoidInjection.injective (Subgroup.fInj (NormalSubgroup.subgroup (trivialSubgroupIsNormal G))) {fzero} {fsucc ()} _
SetoidInjection.injective (Subgroup.fInj (NormalSubgroup.subgroup (trivialSubgroupIsNormal G))) {fsucc ()} {y} _
NormalSubgroup.normal (trivialSubgroupIsNormal {S = S} {_+A_ = _+A_} G) = fzero , transitive (+WellDefined identRight reflexive) invRight
where
open Setoid S
open Group G
open Equivalence eq
improperSubgroupIsNormal : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+A_ : A A A} (G : Group S _+A_) NormalSubgroup G G (identityHom G)
SetoidInjection.wellDefined (Subgroup.fInj (NormalSubgroup.subgroup (improperSubgroupIsNormal G))) {x} {y} x=y = x=y
SetoidInjection.injective (Subgroup.fInj (NormalSubgroup.subgroup (improperSubgroupIsNormal G))) = id
NormalSubgroup.normal (improperSubgroupIsNormal {S = S} {_+A_ = _+A_} G) {g} {h} = ((g +A h) +A inverse g) , reflexive
where
open Group G
open Setoid S
open Equivalence eq

View File

@@ -42,30 +42,3 @@ fourWay+Associative'' {S = S} {_·_ = _·_} G = transitive +Associative (symmetr
open Group G
open Setoid S
open Equivalence eq
{-
quotientHom : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_·A_ : A → A → A} {_·B_ : B → B → B} (G : Group S _·A_) {H : Group T _·B_}{f : A → B} → (fHom : GroupHom G H f) → A → A
quotientHom {S = S} {T = T} {_·A_ = _·A_} {_·B_ = _·B_} G {f = f} fHom a = {!!}
quotientInjection : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_·A_ : A → A → A} {_·B_ : B → B → B} (G : Group S _·A_) {H : Group T _·B_}{f : A → B} → (fHom : GroupHom G H f) → GroupHom (quotientGroup G fHom) G (quotientHom G fHom)
GroupHom.groupHom (quotientInjection {S = S} {T = T} {_·A_ = _·A_} {_·B_ = _·B_} G {f = f} fHom) {x} {y} = {!!}
where
open Setoid S
open Equivalence eq
open Reflexive reflexiveEq
GroupHom.wellDefined (quotientInjection {S = S} {T = T} {_·A_ = _·A_} G {H = H} {f = f} fHom) {x} {y} x~y = {!!}
where
open Group G
open Setoid S
open Setoid T renaming (__ to _T_)
open Equivalence (Setoid.eq S)
open Reflexive reflexiveEq
have : f (x ·A inverse y) T Group.0G H
have = x~y
need : x y
need = {!!}
quotientIsSubgroup : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_·A_ : A → A → A} {_·B_ : B → B → B} {G : Group S _·A_} {H : Group T _·B_}{f : A → B}{fHom : GroupHom G H f} → Subgroup G (quotientGroup G fHom) (quotientInjection G fHom)
quotientIsSubgroup = {!!}
-}

View File

@@ -0,0 +1,37 @@
{-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae
open import Setoids.Setoids
open import Sets.EquivalenceRelations
open import Functions
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Numbers.Naturals.Naturals
open import Sets.FinSet
open import Groups.Definition
module Groups.Subgroups.Examples {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A A A} (G : Group S _+_) where
open import Groups.Subgroups.Definition G
open import Groups.Lemmas G
open Group G
open Setoid S
open Equivalence eq
trivialSubgroupPred : A Set b
trivialSubgroupPred a = (a 0G)
trivialSubgroup : Subgroup trivialSubgroupPred
Subgroup.isSubset trivialSubgroup x=y x=0 = transitive (symmetric x=y) x=0
Subgroup.closedUnderPlus trivialSubgroup x=0 y=0 = transitive (+WellDefined x=0 y=0) identLeft
Subgroup.containsIdentity trivialSubgroup = reflexive
Subgroup.closedUnderInverse trivialSubgroup x=0 = transitive (inverseWellDefined x=0) invIdent
improperSubgroupPred : A Set
improperSubgroupPred a = True
improperSubgroup : Subgroup improperSubgroupPred
Subgroup.isSubset improperSubgroup _ _ = record {}
Subgroup.closedUnderPlus improperSubgroup _ _ = record {}
Subgroup.containsIdentity improperSubgroup = record {}
Subgroup.closedUnderInverse improperSubgroup _ = record {}

View File

@@ -0,0 +1,35 @@
{-# OPTIONS --safe --warning=error --without-K #-}
open import Groups.Groups
open import Groups.Definition
open import Orders
open import Numbers.Integers.Integers
open import Setoids.Setoids
open import LogicalFormulae
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 Groups.QuotientGroup.Definition
open import Groups.Subgroups.Normal.Definition
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
module Groups.Subgroups.Normal.Examples {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A A A} (G : Group S _+_) where
open import Groups.Subgroups.Examples G
open Setoid S
open Equivalence eq
open Group G
trivialSubgroupIsNormal : normalSubgroup G trivialSubgroup
trivialSubgroupIsNormal {g} k=0 = transitive (+WellDefined reflexive (transitive (+WellDefined k=0 reflexive) identLeft)) (invRight {g})
improperSubgroupIsNormal : normalSubgroup G improperSubgroup
improperSubgroupIsNormal _ = record {}

View File

@@ -0,0 +1,15 @@
{-# 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.Groups
open import Groups.Definition
open import Sets.EquivalenceRelations
open import Setoids.Functions.Extension
module Groups.SymmetricGroups.CycleType where

View File

@@ -12,44 +12,45 @@ open import Sets.EquivalenceRelations
open import Setoids.Functions.Extension
module Groups.SymmetricGroups.Definition where
data SymmetryGroupElements {a b : _} {A : Set a} (S : Setoid {a} {b} A) : Set (a b) where
sym : {f : A A} SetoidBijection S S f SymmetryGroupElements S
symmetricSetoid : {a b : _} {A : Set a} (S : Setoid {a} {b} A) Setoid (SymmetryGroupElements S)
Setoid.__ (symmetricSetoid S) (sym {f} bijF) (sym {g} bijG) = ExtensionallyEqual {S = S} {S} (SetoidBijection.wellDefined bijF) (SetoidBijection.wellDefined bijG)
Equivalence.reflexive (Setoid.eq (symmetricSetoid S)) {sym {f} bijF} = extensionallyEqualReflexive S S f (SetoidBijection.wellDefined bijF) (SetoidBijection.wellDefined bijF)
Equivalence.symmetric (Setoid.eq (symmetricSetoid S)) {sym {f} bijF} {sym {g} bijG} f~g = extensionallyEqualSymmetric S S f g (SetoidBijection.wellDefined bijF) (SetoidBijection.wellDefined bijG) f~g
Equivalence.transitive (Setoid.eq (symmetricSetoid S)) {sym {f} bijF} {sym {g} bijG} {sym {h} bijH} f~g g~h = extensionallyEqualTransitive S S f g h (SetoidBijection.wellDefined bijF) (SetoidBijection.wellDefined bijG) (SetoidBijection.wellDefined bijH) f~g g~h
data SymmetryGroupElements {a b : _} {A : Set a} (S : Setoid {a} {b} A) : Set (a b) where
sym : {f : A A} SetoidBijection S S f SymmetryGroupElements S
symmetricGroupOp : {a b : _} {A : Set a} {S : Setoid {a} {b} A} (f g : SymmetryGroupElements S) SymmetryGroupElements S
symmetricGroupOp (sym {f} bijF) (sym {g} bijG) = sym (setoidBijComp bijG bijF)
symmetricSetoid : {a b : _} {A : Set a} (S : Setoid {a} {b} A) Setoid (SymmetryGroupElements S)
Setoid.__ (symmetricSetoid S) (sym {f} bijF) (sym {g} bijG) = ExtensionallyEqual {S = S} {S} (SetoidBijection.wellDefined bijF) (SetoidBijection.wellDefined bijG)
Equivalence.reflexive (Setoid.eq (symmetricSetoid S)) {sym {f} bijF} = extensionallyEqualReflexive S S f (SetoidBijection.wellDefined bijF) (SetoidBijection.wellDefined bijF)
Equivalence.symmetric (Setoid.eq (symmetricSetoid S)) {sym {f} bijF} {sym {g} bijG} f~g = extensionallyEqualSymmetric S S f g (SetoidBijection.wellDefined bijF) (SetoidBijection.wellDefined bijG) f~g
Equivalence.transitive (Setoid.eq (symmetricSetoid S)) {sym {f} bijF} {sym {g} bijG} {sym {h} bijH} f~g g~h = extensionallyEqualTransitive S S f g h (SetoidBijection.wellDefined bijF) (SetoidBijection.wellDefined bijG) (SetoidBijection.wellDefined bijH) f~g g~h
symmetricGroupInv : {a b : _} {A : Set a} (S : Setoid {a} {b} A) SymmetryGroupElements S SymmetryGroupElements S
symmetricGroupInv S (sym {f} bijF) with setoidBijectiveImpliesInvertible bijF
... | record { inverse = inverse ; inverseWellDefined = iwd ; isLeft = isLeft ; isRight = isRight } = sym (setoidInvertibleImpliesBijective (record { fWellDefined = iwd ; inverse = f ; inverseWellDefined = SetoidInjection.wellDefined (SetoidBijection.inj bijF) ; isLeft = λ b isRight b ; isRight = λ b isLeft b }))
symmetricGroupOp : {a b : _} {A : Set a} {S : Setoid {a} {b} A} (f g : SymmetryGroupElements S) SymmetryGroupElements S
symmetricGroupOp (sym {f} bijF) (sym {g} bijG) = sym (setoidBijComp bijG bijF)
symmetricGroupInvIsLeft : {a b : _} {A : Set a} (S : Setoid {a} {b} A) {x : SymmetryGroupElements S} Setoid.__ (symmetricSetoid S) (symmetricGroupOp (symmetricGroupInv S x) x) (sym setoidIdIsBijective)
symmetricGroupInvIsLeft {A = A} S {sym {f = f} fBij} = ans
where
ans : {x : A} Setoid.__ S (SetoidInvertible.inverse (setoidBijectiveImpliesInvertible fBij) (f x)) x
ans {x} with SetoidSurjection.surjective (SetoidBijection.surj fBij) {f x}
ans {x} | a , b = SetoidInjection.injective (SetoidBijection.inj fBij) b
symmetricGroupInv : {a b : _} {A : Set a} (S : Setoid {a} {b} A) SymmetryGroupElements S SymmetryGroupElements S
symmetricGroupInv S (sym {f} bijF) with setoidBijectiveImpliesInvertible bijF
... | record { inverse = inverse ; inverseWellDefined = iwd ; isLeft = isLeft ; isRight = isRight } = sym (setoidInvertibleImpliesBijective (record { fWellDefined = iwd ; inverse = f ; inverseWellDefined = SetoidInjection.wellDefined (SetoidBijection.inj bijF) ; isLeft = λ b isRight b ; isRight = λ b isLeft b }))
symmetricGroupInvIsRight : {a b : _} {A : Set a} (S : Setoid {a} {b} A) {x : SymmetryGroupElements S} Setoid.__ (symmetricSetoid S) (symmetricGroupOp x (symmetricGroupInv S x)) (sym setoidIdIsBijective)
symmetricGroupInvIsRight {A = A} S {sym {f = f} fBij} = ans
where
ans : {x : A} Setoid.__ S (f (SetoidInvertible.inverse (setoidBijectiveImpliesInvertible fBij) x)) x
ans {x} with SetoidSurjection.surjective (SetoidBijection.surj fBij) {x}
ans {x} | a , b = b
symmetricGroupInvIsLeft : {a b : _} {A : Set a} (S : Setoid {a} {b} A) {x : SymmetryGroupElements S} Setoid.__ (symmetricSetoid S) (symmetricGroupOp (symmetricGroupInv S x) x) (sym setoidIdIsBijective)
symmetricGroupInvIsLeft {A = A} S {sym {f = f} fBij} = ans
where
ans : {x : A} Setoid.__ S (SetoidInvertible.inverse (setoidBijectiveImpliesInvertible fBij) (f x)) x
ans {x} with SetoidSurjection.surjective (SetoidBijection.surj fBij) {f x}
ans {x} | a , b = SetoidInjection.injective (SetoidBijection.inj fBij) b
symmetricGroup : {a b : _} {A : Set a} (S : Setoid {a} {b} A) Group (symmetricSetoid S) (symmetricGroupOp {A = A})
Group.+WellDefined (symmetricGroup A) {sym {m} bijM} {sym {n} bijN} {sym {x} bijX} {sym {y} bijY} m~x n~y = transitive m~x (SetoidBijection.wellDefined bijX n~y)
where
open Equivalence (Setoid.eq A)
Group.0G (symmetricGroup A) = sym setoidIdIsBijective
Group.inverse (symmetricGroup S) = symmetricGroupInv S
Group.+Associative (symmetricGroup A) {sym {f} bijF} {sym {g} bijG} {sym {h} bijH} = Equivalence.reflexive (Setoid.eq A)
Group.identRight (symmetricGroup A) {sym {f} bijF} = Equivalence.reflexive (Setoid.eq A)
Group.identLeft (symmetricGroup A) {sym {f} bijF} = Equivalence.reflexive (Setoid.eq A)
Group.invLeft (symmetricGroup S) {x} = symmetricGroupInvIsLeft S {x}
Group.invRight (symmetricGroup S) {x} = symmetricGroupInvIsRight S {x}
symmetricGroupInvIsRight : {a b : _} {A : Set a} (S : Setoid {a} {b} A) {x : SymmetryGroupElements S} Setoid.__ (symmetricSetoid S) (symmetricGroupOp x (symmetricGroupInv S x)) (sym setoidIdIsBijective)
symmetricGroupInvIsRight {A = A} S {sym {f = f} fBij} = ans
where
ans : {x : A} Setoid.__ S (f (SetoidInvertible.inverse (setoidBijectiveImpliesInvertible fBij) x)) x
ans {x} with SetoidSurjection.surjective (SetoidBijection.surj fBij) {x}
ans {x} | a , b = b
symmetricGroup : {a b : _} {A : Set a} (S : Setoid {a} {b} A) Group (symmetricSetoid S) (symmetricGroupOp {A = A})
Group.+WellDefined (symmetricGroup A) {sym {m} bijM} {sym {n} bijN} {sym {x} bijX} {sym {y} bijY} m~x n~y = transitive m~x (SetoidBijection.wellDefined bijX n~y)
where
open Equivalence (Setoid.eq A)
Group.0G (symmetricGroup A) = sym setoidIdIsBijective
Group.inverse (symmetricGroup S) = symmetricGroupInv S
Group.+Associative (symmetricGroup A) {sym {f} bijF} {sym {g} bijG} {sym {h} bijH} = Equivalence.reflexive (Setoid.eq A)
Group.identRight (symmetricGroup A) {sym {f} bijF} = Equivalence.reflexive (Setoid.eq A)
Group.identLeft (symmetricGroup A) {sym {f} bijF} = Equivalence.reflexive (Setoid.eq A)
Group.invLeft (symmetricGroup S) {x} = symmetricGroupInvIsLeft S {x}
Group.invRight (symmetricGroup S) {x} = symmetricGroupInvIsRight S {x}

View File

@@ -0,0 +1,31 @@
{-# 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.Semiring
open import Sets.FinSet
open import Groups.Groups
open import Groups.Definition
open import Sets.EquivalenceRelations
open import Setoids.Functions.Extension
open import Groups.SymmetricGroups.Definition
module Groups.SymmetricGroups.Finite.Definition where
snSet : (n : ) Set
snSet n = SymmetryGroupElements (reflSetoid (FinSet n))
snSetoid : (n : ) Setoid (snSet n)
snSetoid n = symmetricSetoid (reflSetoid (FinSet n))
SymmetricGroupN : (n : ) Group (snSetoid n) (symmetricGroupOp)
SymmetricGroupN n = symmetricGroup (reflSetoid (FinSet n))
record Cycles {n : } (s : snSet n) : Set where
field
cycles : {n : } (s : snSet n) Cycles s
cycles s = {!!}

View File

@@ -15,7 +15,6 @@ data False' {a : _} : Set a where
record True : Set where
record True' {a : _} : Set a where
record Unit : Set where
infix 10 _||_
data _||_ {a b} (A : Set a) (B : Set b) : Set (a b) where

View File

@@ -0,0 +1,107 @@
{-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae
open import Numbers.Naturals.Semiring
open import Numbers.Naturals.Order
open import Numbers.Naturals.Naturals
open import Numbers.Naturals.Multiplication
open import Numbers.Integers.Order
open import Numbers.Integers.RingStructure.Ring
open import Numbers.Integers.RingStructure.IntegralDomain
open import Semirings.Definition
open import Groups.Definition
open import Rings.Definition
open import Setoids.Setoids
open import Rings.EuclideanDomains.Definition
open import Orders
open import Numbers.Naturals.EuclideanAlgorithm
module Numbers.Integers.RingStructure.EuclideanDomain where
private
norm :
norm (nonneg x) = x
norm (negSucc x) = succ x
norm' : {a : } (a!=0 : (a nonneg 0) False)
norm' {a} _ = norm a
multiplyIncreasesSuccLemma : (a b : ) succ (succ a) <N (succ (succ a)) *N (succ (succ b))
multiplyIncreasesSuccLemma a b with multiplyIncreases (succ (succ b)) (succ (succ a)) (le b (applyEquality succ (Semiring.commutative Semiring b 1))) (le (succ a) (applyEquality (λ i succ (succ i)) (Semiring.sumZeroRight Semiring a)))
... | bl rewrite multiplicationNIsCommutative (succ (succ b)) (succ (succ a)) = bl
multiplyIncreasesSucc : (a b : ) succ a ≤N (succ a) *N (succ b)
multiplyIncreasesSucc zero zero = inr refl
multiplyIncreasesSucc zero (succ b) = inl (le (b +N 0) (Semiring.commutative Semiring (succ (b +N 0)) 1))
multiplyIncreasesSucc (succ a) zero = inr (applyEquality (λ i succ (succ i)) (equalityCommutative (productWithOneRight a)))
multiplyIncreasesSucc (succ a) (succ b) = inl (multiplyIncreasesSuccLemma a b)
multiplyIncreasesSuccLemma' : (a b : ) succ (succ a) <N succ ((succ (succ a)) *N succ b) +N succ a
multiplyIncreasesSuccLemma' a b = succPreservesInequality (le (b +N succ (b +N a *N succ b)) refl)
multiplyIncreasesSucc' : (a b : ) succ a ≤N succ ((b +N a *N b) +N a)
multiplyIncreasesSucc' zero zero = inr refl
multiplyIncreasesSucc' zero (succ b) = inl (le b (applyEquality succ (transitivity (Semiring.commutative Semiring b 1) (transitivity (equalityCommutative (Semiring.sumZeroRight Semiring (succ b))) (equalityCommutative (Semiring.sumZeroRight Semiring _))))))
multiplyIncreasesSucc' (succ a) zero rewrite multiplicationNIsCommutative a zero = inr refl
multiplyIncreasesSucc' (succ a) (succ b) = inl (multiplyIncreasesSuccLemma' a b)
normSize : {a b : } (a!=0 : (a nonneg 0) False) (b!=0 : (b nonneg 0) False) (c : ) b (a *Z c) (norm a) ≤N (norm b)
normSize {nonneg a} {nonneg b} a!=0 b!=0 (nonneg zero) b=ac rewrite nonnegInjective b=ac | multiplicationNIsCommutative a 0 = exFalso (b!=0 refl)
normSize {nonneg a} {nonneg b} a!=0 b!=0 (nonneg (succ 0)) b=ac rewrite nonnegInjective b=ac | multiplicationNIsCommutative a 1 = inr (equalityCommutative (Semiring.sumZeroRight Semiring a))
normSize {nonneg 0} {nonneg b} a!=0 b!=0 (nonneg (succ (succ c))) b=ac = exFalso (a!=0 refl)
normSize {nonneg (succ a)} {nonneg (succ b)} a!=0 b!=0 (nonneg (succ (succ c))) b=ac rewrite nonnegInjective b=ac = multiplyIncreasesSucc a (succ c)
normSize {nonneg 0} {nonneg b} a!=0 b!=0 (negSucc c) bad = exFalso (a!=0 refl)
normSize {nonneg (succ a)} {nonneg b} a!=0 b!=0 (negSucc c) ()
normSize {nonneg a} {negSucc b} a!=0 b!=0 (nonneg c) ()
normSize {nonneg (succ a)} {negSucc b} a!=0 b!=0 (negSucc c) pr rewrite negSuccInjective pr = multiplyIncreasesSucc' a c
normSize {negSucc a} {nonneg zero} _ b!=0 c pr = exFalso (b!=0 refl)
normSize {negSucc a} {nonneg (succ b)} _ _ (nonneg zero) ()
normSize {negSucc a} {nonneg (succ b)} _ _ (nonneg (succ x)) ()
normSize {negSucc a} {nonneg (succ b)} _ _ (negSucc c) pr rewrite nonnegInjective pr = multiplyIncreasesSucc a c
normSize {negSucc a} {negSucc b} _ _ (nonneg (succ c)) pr rewrite negSuccInjective pr | multiplicationNIsCommutative c a | Semiring.commutative Semiring (a +N a *N c) c | Semiring.+Associative Semiring c a (a *N c) | Semiring.commutative Semiring c a | equalityCommutative (Semiring.+Associative Semiring a c (a *N c)) | Semiring.commutative Semiring a (c +N a *N c) = multiplyIncreasesSucc' a c
divAlg : {a b : } (a!=0 : (a nonneg 0) False) (b!=0 : (b nonneg 0) False) DivisionAlgorithmResult Ring norm' a!=0 b!=0
divAlg {nonneg zero} {b} a!=0 b!=0 = exFalso (a!=0 refl)
divAlg {nonneg (succ a)} {nonneg zero} a!=0 b!=0 = exFalso (b!=0 refl)
divAlg {nonneg (succ a)} {nonneg (succ b)} a!=0 b!=0 with divisionAlg (succ b) (succ a)
divAlg {nonneg (succ a)} {nonneg (succ b)} a!=0 b!=0 | record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = inl remIsSmall ; quotSmall = quotSmall } = record { quotient = nonneg quot ; rem = nonneg rem ; remSmall = u ; divAlg = transitivity (applyEquality nonneg (equalityCommutative pr)) t }
where
t : nonneg ((quot +N b *N quot) +N rem) nonneg (quot *N succ b) +Z nonneg rem
t rewrite multiplicationNIsCommutative (succ b) quot = +Zinherits (quot *N succ b) rem
u : (nonneg rem nonneg 0) || Sg ((nonneg rem nonneg 0) False) (λ pr rem <N succ b)
u with TotalOrder.totality Order (nonneg 0) (nonneg rem)
u | inl (inl 0<rem) = inr ((λ p TotalOrder.irreflexive Order (TotalOrder.<WellDefined Order refl p 0<rem)) , remIsSmall)
u | inr 0=rem rewrite 0=rem = inl refl
divAlg {nonneg (succ a)} {negSucc b} _ _ with divisionAlg (succ b) (succ a)
divAlg {nonneg (succ a)} {negSucc b} _ _ | record { quot = zero ; rem = rem ; pr = pr ; remIsSmall = inl remIsSmall ; quotSmall = inl quotSmall } rewrite multiplicationNIsCommutative b 0 = record { quotient = nonneg 0 ; rem = nonneg rem ; remSmall = inr (p , remIsSmall) ; divAlg = applyEquality nonneg (equalityCommutative pr) }
where
p : (nonneg rem nonneg 0) False
p pr2 rewrite pr = naughtE (equalityCommutative (nonnegInjective pr2))
divAlg {nonneg (succ a)} {negSucc b} _ _ | record { quot = succ quot ; rem = zero ; pr = pr ; remIsSmall = inl remIsSmall ; quotSmall = inl quotSmall } = record { quotient = negSucc quot ; rem = nonneg zero ; remSmall = inl refl ; divAlg = applyEquality nonneg (transitivity (equalityCommutative pr) (applyEquality (λ i i +N 0) (multiplicationNIsCommutative (succ b) (succ quot)))) }
divAlg {nonneg (succ a)} {negSucc b} _ _ | record { quot = succ quot ; rem = succ rem ; pr = pr ; remIsSmall = inl remIsSmall ; quotSmall = inl quotSmall } = record { quotient = negSucc quot ; rem = nonneg (succ rem) ; remSmall = inr ((λ ()) , remIsSmall) ; divAlg = applyEquality nonneg (transitivity (equalityCommutative pr) (applyEquality (λ i i +N succ rem) (multiplicationNIsCommutative (succ b) (succ quot)))) }
divAlg {negSucc a} {nonneg zero} _ b!=0 = exFalso (b!=0 refl)
divAlg {negSucc a} {nonneg (succ b)} _ _ with divisionAlg (succ b) (succ a)
divAlg {negSucc a} {nonneg (succ b)} _ _ | record { quot = zero ; rem = zero ; pr = pr ; remIsSmall = inl remIsSmall ; quotSmall = quotSmall } = exFalso (naughtE (transitivity (equalityCommutative (transitivity (Semiring.sumZeroRight Semiring _) (multiplicationNIsCommutative b 0))) pr))
divAlg {negSucc a} {nonneg (succ b)} _ _ | record { quot = succ quot ; rem = zero ; pr = pr ; remIsSmall = inl remIsSmall ; quotSmall = quotSmall } = record { quotient = negSucc quot ; rem = nonneg 0 ; remSmall = inl refl ; divAlg = applyEquality negSucc (succInjective (transitivity (equalityCommutative pr) t)) }
where
t : succ (quot +N b *N succ quot) +N 0 succ ((quot +N b *N quot) +N b)
t rewrite Semiring.sumZeroRight Semiring (succ (quot +N b *N succ quot)) | Semiring.commutative Semiring (quot +N b *N quot) b | Semiring.+Associative Semiring b quot (b *N quot) | Semiring.commutative Semiring b quot | equalityCommutative (Semiring.+Associative Semiring quot b (b *N quot)) | multiplicationNIsCommutative b (succ quot) | multiplicationNIsCommutative quot b = refl
divAlg {negSucc a} {nonneg (succ b)} _ _ | record { quot = zero ; rem = succ rem ; pr = pr ; remIsSmall = inl remIsSmall ; quotSmall = quotSmall } rewrite multiplicationNIsCommutative b 0 | equalityCommutative (succInjective pr) = record { quotient = nonneg zero ; rem = negSucc rem ; remSmall = inr (((λ ()) , remIsSmall)) ; divAlg = refl }
divAlg {negSucc a} {nonneg (succ b)} _ _ | record { quot = succ quot ; rem = succ rem ; pr = pr ; remIsSmall = inl remIsSmall ; quotSmall = quotSmall } = record { quotient = negSucc quot ; rem = negSucc rem ; remSmall = inr ((λ ()) , remIsSmall) ; divAlg = applyEquality negSucc (succInjective (transitivity (equalityCommutative pr) t)) }
where
t : succ b *N succ quot +N succ rem succ (succ (succ b *N quot +N b) +N rem)
t rewrite Semiring.commutative Semiring ((quot +N b *N quot) +N b) rem | Semiring.commutative Semiring (succ rem) ((quot +N b *N quot) +N b) | multiplicationNIsCommutative b (succ quot) | equalityCommutative (Semiring.+Associative Semiring quot (b *N quot) b) | Semiring.commutative Semiring b (quot *N b) | multiplicationNIsCommutative b quot = refl
divAlg {negSucc a} {negSucc b} _ _ with divisionAlg (succ b) (succ a)
divAlg {negSucc a} {negSucc b} _ _ | record { quot = zero ; rem = zero ; pr = pr ; remIsSmall = inl remIsSmall ; quotSmall = quotSmall } = exFalso (naughtE (transitivity (equalityCommutative (transitivity (Semiring.sumZeroRight Semiring (b *N 0)) (multiplicationNIsCommutative b 0))) pr))
divAlg {negSucc a} {negSucc b} _ _ | record { quot = zero ; rem = succ rem ; pr = pr ; remIsSmall = inl remIsSmall ; quotSmall = quotSmall } rewrite multiplicationNIsCommutative b 0 | equalityCommutative (succInjective pr) = record { quotient = nonneg 0 ; rem = negSucc rem ; remSmall = inr ((λ ()) , remIsSmall) ; divAlg = refl }
divAlg {negSucc a} {negSucc b} _ _ | record { quot = succ quot ; rem = zero ; pr = pr ; remIsSmall = inl remIsSmall ; quotSmall = quotSmall } rewrite Semiring.sumZeroRight Semiring (quot +N b *N succ quot) | Semiring.commutative Semiring b (succ quot) = record { quotient = nonneg (succ quot) ; rem = nonneg 0 ; remSmall = inl refl ; divAlg = applyEquality negSucc (succInjective (transitivity (equalityCommutative pr) t)) }
where
t : succ (quot +N b *N succ quot) succ ((b +N quot *N b) +N quot)
t rewrite Semiring.commutative Semiring quot (b *N succ quot) | multiplicationNIsCommutative b (succ quot) = refl
divAlg {negSucc a} {negSucc b} _ _ | record { quot = succ quot ; rem = succ rem ; pr = pr ; remIsSmall = inl remIsSmall ; quotSmall = quotSmall } rewrite multiplicationNIsCommutative b (succ quot) | Semiring.commutative Semiring quot (b +N quot *N b) | Semiring.commutative Semiring ((b +N quot *N b) +N quot) (succ rem) | Semiring.commutative Semiring rem ((b +N quot *N b) +N quot) = record { quotient = nonneg (succ quot) ; rem = negSucc rem ; remSmall = inr (((λ ()) , remIsSmall)) ; divAlg = applyEquality negSucc (equalityCommutative (succInjective pr)) }
EuclideanDomain : EuclideanDomain Ring
EuclideanDomain.isIntegralDomain EuclideanDomain = IntDom
EuclideanDomain.norm EuclideanDomain = norm'
EuclideanDomain.normSize EuclideanDomain = normSize
EuclideanDomain.divisionAlg EuclideanDomain {a} {b} a!=0 b!=0 = divAlg {a} {b} a!=0 b!=0

View File

@@ -0,0 +1,358 @@
{-# OPTIONS --warning=error --safe --without-K #-}
open import LogicalFormulae
open import Numbers.Naturals.Semiring
open import Numbers.Naturals.Naturals
open import Numbers.Naturals.Order
open import Numbers.Naturals.Order.Lemmas
open import Numbers.Naturals.Order.WellFounded
open import WellFoundedInduction
open import KeyValue.KeyValue
open import Orders
open import Vectors
open import Maybe
open import Semirings.Definition
module Numbers.Naturals.EuclideanAlgorithm where
open TotalOrder TotalOrder
open Semiring Semiring
record divisionAlgResult (a : ) (b : ) : Set where
field
quot :
rem :
pr : a *N quot +N rem b
remIsSmall : (rem <N a) || (a 0)
quotSmall : (0 <N a) || ((0 a) && (quot 0))
divAlgLessLemma : (a b : ) (0 <N a) (r : divisionAlgResult a b) (divisionAlgResult.quot r 0) || (divisionAlgResult.rem r <N b)
divAlgLessLemma zero b pr _ = exFalso (TotalOrder.irreflexive TotalOrder pr)
divAlgLessLemma (succ a) b _ record { quot = zero ; rem = a%b ; pr = pr ; remIsSmall = remIsSmall } = inl refl
divAlgLessLemma (succ a) b _ record { quot = (succ a/b) ; rem = a%b ; pr = pr ; remIsSmall = remIsSmall } = inr record { x = a/b +N a *N succ (a/b) ; proof = pr }
modUniqueLemma : {rem1 rem2 a : } (quot1 quot2 : ) rem1 <N a rem2 <N a a *N quot1 +N rem1 a *N quot2 +N rem2 rem1 rem2
modUniqueLemma {rem1} {rem2} {a} zero zero rem1<a rem2<a pr rewrite Semiring.productZeroRight Semiring a = pr
modUniqueLemma {rem1} {rem2} {a} zero (succ quot2) rem1<a rem2<a pr rewrite Semiring.productZeroRight Semiring a | pr | multiplicationNIsCommutative a (succ quot2) | equalityCommutative (Semiring.+Associative Semiring a (quot2 *N a) rem2) = exFalso (cannotAddAndEnlarge' {a} {quot2 *N a +N rem2} rem1<a)
modUniqueLemma {rem1} {rem2} {a} (succ quot1) zero rem1<a rem2<a pr rewrite Semiring.productZeroRight Semiring a | equalityCommutative pr | multiplicationNIsCommutative a (succ quot1) | equalityCommutative (Semiring.+Associative Semiring a (quot1 *N a) rem1) = exFalso (cannotAddAndEnlarge' {a} {quot1 *N a +N rem1} rem2<a)
modUniqueLemma {rem1} {rem2} {a} (succ quot1) (succ quot2) rem1<a rem2<a pr rewrite multiplicationNIsCommutative a (succ quot1) | multiplicationNIsCommutative a (succ quot2) | equalityCommutative (Semiring.+Associative Semiring a (quot1 *N a) rem1) | equalityCommutative (Semiring.+Associative Semiring a (quot2 *N a) rem2) = modUniqueLemma {rem1} {rem2} {a} quot1 quot2 rem1<a rem2<a (go {a}{quot1}{rem1}{quot2}{rem2} pr)
where
go : {a quot1 rem1 quot2 rem2 : } (a +N (quot1 *N a +N rem1) a +N (quot2 *N a +N rem2)) a *N quot1 +N rem1 a *N quot2 +N rem2
go {a} {quot1} {rem1} {quot2} {rem2} pr rewrite multiplicationNIsCommutative quot1 a | multiplicationNIsCommutative quot2 a = canSubtractFromEqualityLeft {a} pr
modIsUnique : {a b : } (div1 div2 : divisionAlgResult a b) divisionAlgResult.rem div1 divisionAlgResult.rem div2
modIsUnique {zero} {b} record { quot = quot1 ; rem = rem1 ; pr = pr1 ; remIsSmall = remIsSmall1 } record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = remIsSmall } = transitivity pr1 (equalityCommutative pr)
modIsUnique {succ a} {b} record { quot = quot1 ; rem = rem1 ; pr = pr1 ; remIsSmall = (inl y) } record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = (inl x) } = modUniqueLemma {rem1} {rem} {succ a} quot1 quot y x (transitivity pr1 (equalityCommutative pr))
modIsUnique {succ a} {b} record { quot = quot1 ; rem = rem1 ; pr = pr1 ; remIsSmall = (inr ()) } record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = (inl x) }
modIsUnique {succ a} {b} record { quot = quot1 ; rem = rem1 ; pr = pr1 ; remIsSmall = remIsSmall1 } record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = (inr ()) }
transferAddition : (a b c : ) (a +N b) +N c (a +N c) +N b
transferAddition a b c rewrite equalityCommutative (Semiring.+Associative Semiring a b c) = p a b c
where
p : (a b c : ) a +N (b +N c) (a +N c) +N b
p a b c = transitivity (applyEquality (a +N_) (Semiring.commutative Semiring b c)) (Semiring.+Associative Semiring a c b)
divisionAlgLemma : (x b : ) x *N zero +N b b
divisionAlgLemma x b rewrite (Semiring.productZeroRight Semiring x) = refl
divisionAlgLemma2 : (x b : ) (x b) x *N succ zero +N zero b
divisionAlgLemma2 x b pr rewrite (Semiring.productOneRight Semiring x) = equalityCommutative (transitivity (equalityCommutative pr) (equalityCommutative (Semiring.sumZeroRight Semiring x)))
divisionAlgLemma3 : {a x : } (p : succ a <N succ x) (subtractionNResult.result (-N (inl p))) <N (succ x)
divisionAlgLemma3 {a} {x} p = -NIsDecreasing {a} {succ x} p
divisionAlgLemma4 : (p a q : ) ((p +N a *N p) +N q) +N succ a succ ((p +N a *N succ p) +N q)
divisionAlgLemma4 p a q = ans
where
r : ((p +N a *N p) +N q) +N succ a succ (((p +N a *N p) +N q) +N a)
ans : ((p +N a *N p) +N q) +N succ a succ ((p +N a *N succ p) +N q)
s : ((p +N a *N p) +N q) +N a (p +N a *N succ p) +N q
t : (p +N a *N p) +N a p +N a *N succ p
s = transitivity (transferAddition (p +N a *N p) q a) (applyEquality (λ i i +N q) t)
ans = identityOfIndiscernablesRight _≡_ r (applyEquality succ s)
r = succExtracts ((p +N a *N p) +N q) a
t = transitivity (equalityCommutative (Semiring.+Associative Semiring p (a *N p) a)) (applyEquality (λ n p +N n) (equalityCommutative (transitivity (multiplicationNIsCommutative a (succ p)) (transitivity (Semiring.commutative Semiring a _) (applyEquality (_+N a) (multiplicationNIsCommutative p _))))))
divisionAlg : (a : ) (b : ) divisionAlgResult a b
divisionAlg zero = λ b record { quot = zero ; rem = b ; pr = refl ; remIsSmall = inr refl ; quotSmall = inr (record { fst = refl ; snd = refl }) }
divisionAlg (succ a) = rec <NWellfounded (λ n divisionAlgResult (succ a) n) go
where
go : (x : ) (indHyp : (y : ) (y<x : y <N x) divisionAlgResult (succ a) y)
divisionAlgResult (succ a) x
go zero prop = record { quot = zero ; rem = zero ; pr = divisionAlgLemma (succ a) zero ; remIsSmall = inl (succIsPositive a) ; quotSmall = inl (succIsPositive a) }
go (succ x) indHyp with totality (succ a) (succ x)
go (succ x) indHyp | inl (inl sa<sx) with indHyp (subtractionNResult.result (-N (inl sa<sx))) (divisionAlgLemma3 sa<sx)
... | record { quot = prevQuot ; rem = prevRem ; pr = prevPr ; remIsSmall = smallRem } = p
where
p : divisionAlgResult (succ a) (succ x)
addedA : (succ a *N prevQuot +N prevRem) +N (succ a) subtractionNResult.result (-N (inl sa<sx)) +N (succ a)
addedA' : (succ a *N prevQuot +N prevRem) +N succ a succ x
addedA'' : (succ a *N succ prevQuot) +N prevRem succ x
addedA''' : succ ((prevQuot +N a *N succ prevQuot) +N prevRem) succ x
addedA''' = identityOfIndiscernablesLeft _≡_ addedA'' refl
p = record { quot = succ prevQuot ; rem = prevRem ; pr = addedA''' ; remIsSmall = smallRem ; quotSmall = inl (succIsPositive a) }
addedA = applyEquality (λ n n +N succ a) prevPr
addedA' = identityOfIndiscernablesRight _≡_ addedA (addMinus {succ a} {succ x} (inl sa<sx))
addedA'' = identityOfIndiscernablesLeft _≡_ addedA' (divisionAlgLemma4 prevQuot a prevRem)
go (succ x) indHyp | inr (sa=sx) = record { quot = succ zero ; rem = zero ; pr = divisionAlgLemma2 (succ a) (succ x) sa=sx ; remIsSmall = inl (succIsPositive a) ; quotSmall = inl (succIsPositive a) }
go (succ x) indHyp | inl (inr (sx<sa)) = record { quot = zero ; rem = succ x ; pr = divisionAlgLemma (succ a) (succ x) ; remIsSmall = inl sx<sa ; quotSmall = inl (succIsPositive a) }
data __ : Set where
divides : {a b : } (res : divisionAlgResult a b) divisionAlgResult.rem res zero a b
zeroDividesNothing : (a : ) zero succ a False
zeroDividesNothing a (divides record { quot = quot ; rem = rem ; pr = pr } x) = naughtE p
where
p : zero succ a
p = transitivity (equalityCommutative x) pr
record hcfData (a b : ) : Set where
field
c :
c|a : c a
c|b : c b
hcf : x x a x b x c
positiveTimes : {a b : } (succ a *N succ b <N succ a) False
positiveTimes {a} {b} pr = zeroNeverGreater f'
where
g : succ a *N succ b <N succ a *N succ 0
g rewrite multiplicationNIsCommutative a 1 | Semiring.commutative Semiring a 0 = pr
f : succ b <N succ 0
f = cancelInequalityLeft {succ a} {succ b} g
f' : b <N 0
f' = canRemoveSuccFrom<N f
biggerThanCantDivideLemma : {a b : } (a <N b) (b a) a 0
biggerThanCantDivideLemma {zero} {b} a<b b|a = refl
biggerThanCantDivideLemma {succ a} {zero} a<b (divides record { quot = quot ; rem = .0 ; pr = pr ; remIsSmall = remIsSmall ; quotSmall = (inl (le x ())) } refl)
biggerThanCantDivideLemma {succ a} {zero} a<b (divides record { quot = quot ; rem = .0 ; pr = () ; remIsSmall = remIsSmall ; quotSmall = (inr x) } refl)
biggerThanCantDivideLemma {succ a} {succ b} a<b (divides record { quot = zero ; rem = .0 ; pr = pr ; remIsSmall = remIsSmall ; quotSmall = quotSmall } refl) rewrite Semiring.commutative Semiring (b *N zero) 0 | multiplicationNIsCommutative b 0 = exFalso (naughtE pr)
biggerThanCantDivideLemma {succ a} {succ b} a<b (divides record { quot = (succ quot) ; rem = .0 ; pr = pr ; remIsSmall = remIsSmall ; quotSmall = quotSmall } refl) rewrite Semiring.commutative Semiring (quot +N b *N succ quot) 0 | equalityCommutative pr = exFalso (positiveTimes {b} {quot} a<b)
biggerThanCantDivide : {a b : } (x : ) (TotalOrder.max TotalOrder a b) <N x (x a) (x b) (a 0) && (b 0)
biggerThanCantDivide {a} {b} x pr x|a x|b with totality a b
biggerThanCantDivide {a} {b} x pr x|a x|b | inl (inl a<b) = exFalso (zeroNeverGreater f')
where
f : b 0
f = biggerThanCantDivideLemma pr x|b
f' : a <N 0
f' rewrite equalityCommutative f = a<b
biggerThanCantDivide {a} {b} x pr x|a x|b | inl (inr b<a) = exFalso (zeroNeverGreater f')
where
f : a 0
f = biggerThanCantDivideLemma pr x|a
f' : b <N 0
f' rewrite equalityCommutative f = b<a
biggerThanCantDivide {a} {b} x pr x|a x|b | inr a=b = (transitivity a=b f ,, f)
where
f : b 0
f = biggerThanCantDivideLemma pr x|b
aDivA : (a : ) a a
aDivA zero = divides (record { quot = 0 ; rem = 0 ; pr = equalityCommutative (oneTimesPlusZero zero) ; remIsSmall = inr refl ; quotSmall = inr (refl ,, refl) }) refl
aDivA (succ a) = divides (record { quot = 1 ; rem = 0 ; pr = equalityCommutative (oneTimesPlusZero (succ a)) ; remIsSmall = inl (succIsPositive a) ; quotSmall = inl (succIsPositive a) }) refl
aDivZero : (a : ) a zero
aDivZero zero = aDivA zero
aDivZero (succ a) = divides (record { quot = zero ; rem = zero ; pr = lemma (succ a) ; remIsSmall = inl (succIsPositive a) ; quotSmall = inl (succIsPositive a) }) refl
where
lemma : (b : ) b *N zero +N zero zero
lemma b rewrite (Semiring.sumZeroRight Semiring (b *N zero)) = Semiring.productZeroRight Semiring b
record extendedHcf (a b : ) : Set where
field
hcf : hcfData a b
c :
c = hcfData.c hcf
field
extended1 :
extended2 :
extendedProof : (a *N extended1 b *N extended2 +N c) || (a *N extended1 +N c b *N extended2)
divEqualityLemma1 : {a b c : } b zero b *N c +N 0 a a b
divEqualityLemma1 {a} {.0} {c} refl pr = equalityCommutative pr
divEquality : {a b : } a b b a a b
divEquality {a} {b} (divides record { quot = quotAB ; rem = .0 ; pr = prAB ; remIsSmall = _ ; quotSmall = quotSmallAB } refl) (divides record { quot = quot ; rem = .0 ; pr = pr ; remIsSmall = _ ; quotSmall = (inl x) } refl) rewrite Semiring.commutative Semiring (b *N quot) 0 | Semiring.commutative Semiring (a *N quotAB) 0 | equalityCommutative pr | equalityCommutative (Semiring.*Associative Semiring b quot quotAB) = res
where
lem : {b r : } b *N r b (0 <N b) r 1
lem {zero} {r} pr ()
lem {succ b} {zero} pr _ rewrite multiplicationNIsCommutative b 0 = exFalso (naughtE pr)
lem {succ b} {succ zero} pr _ = refl
lem {succ b} {succ (succ r)} pr _ rewrite multiplicationNIsCommutative b (succ (succ r)) | Semiring.commutative Semiring (succ r) (b +N (b +N r *N b)) | equalityCommutative (Semiring.+Associative Semiring b (b +N r *N b) (succ r)) | Semiring.commutative Semiring (b +N r *N b) (succ r) = exFalso (cannotAddAndEnlarge'' {succ b} pr)
p : quot *N quotAB 1
p = lem prAB x
q : quot 1
q = _&&_.fst (productOneImpliesOperandsOne p)
res : b *N quot b
res rewrite q | multiplicationNIsCommutative b 1 | Semiring.commutative Semiring b 0 = refl
divEquality {.0} {.0} (divides record { quot = quotAB ; rem = .0 ; pr = prAB ; remIsSmall = _ ; quotSmall = quotSmallAB } refl) (divides record { quot = quot ; rem = .0 ; pr = refl ; remIsSmall = _ ; quotSmall = (inr (refl ,, snd)) } refl) = refl
hcfWelldefined : {a b : } (ab : hcfData a b) (ab' : hcfData a b) (hcfData.c ab hcfData.c ab')
hcfWelldefined {a} {b} record { c = c ; c|a = c|a ; c|b = c|b ; hcf = hcf } record { c = c' ; c|a = c|a' ; c|b = c|b' ; hcf = hcf' } with hcf c' c|a' c|b'
... | c'DivC with hcf' c c|a c|b
... | cDivC' = divEquality cDivC' c'DivC
reverseHCF : {a b : } (ab : extendedHcf a b) extendedHcf b a
reverseHCF {a} {b} record { hcf = record { c = c ; c|a = c|a ; c|b = c|b ; hcf = hcf } ; extended1 = extended1 ; extended2 = extended2 ; extendedProof = (inl x) } = record { hcf = record { c = c ; c|a = c|b ; c|b = c|a ; hcf = λ x z z₁ hcf x z₁ z } ; extended1 = extended2 ; extended2 = extended1 ; extendedProof = inr (equalityCommutative x) }
reverseHCF {a} {b} record { hcf = record { c = c ; c|a = c|a ; c|b = c|b ; hcf = hcf } ; extended1 = extended1 ; extended2 = extended2 ; extendedProof = (inr x) } = record { hcf = record { c = c ; c|a = c|b ; c|b = c|a ; hcf = λ x z z₁ hcf x z₁ z } ; extended1 = extended2 ; extended2 = extended1 ; extendedProof = inl (equalityCommutative x) }
oneDivN : (a : ) 1 a
oneDivN a = divides (record { quot = a ; rem = zero ; pr = pr ; remIsSmall = inl (succIsPositive zero) ; quotSmall = inl (le zero refl) }) refl
where
pr : (a +N zero) +N zero a
pr rewrite Semiring.sumZeroRight Semiring (a +N zero) = Semiring.sumZeroRight Semiring a
hcfZero : (a : ) extendedHcf zero a
hcfZero a = record { hcf = record { c = a ; c|a = aDivZero a ; c|b = aDivA a ; hcf = λ _ _ p p } ; extended1 = 0 ; extended2 = 1 ; extendedProof = inr (equalityCommutative (Semiring.productOneRight Semiring a))}
hcfOne : (a : ) extendedHcf 1 a
hcfOne a = record { hcf = record { c = 1 ; c|a = aDivA 1 ; c|b = oneDivN a ; hcf = λ _ z _ z } ; extended1 = 1 ; extended2 = 0 ; extendedProof = inl g }
where
g : 1 a *N 0 +N 1
g rewrite multiplicationNIsCommutative a 0 = refl
zeroIsValidRem : (a : ) (0 <N a) || (a 0)
zeroIsValidRem zero = inr refl
zeroIsValidRem (succ a) = inl (succIsPositive a)
dividesBothImpliesDividesSum : {a x y : } a x a y a (x +N y)
dividesBothImpliesDividesSum {a} {x} {y} (divides record { quot = xDivA ; rem = .0 ; pr = prA ; quotSmall = qsm1 } refl) (divides record { quot = quot ; rem = .0 ; pr = pr ; quotSmall = qsm2 } refl) = divides (record { quot = xDivA +N quot ; rem = 0 ; pr = go {a} {x} {y} {xDivA} {quot} pr prA ; remIsSmall = zeroIsValidRem a ; quotSmall = (quotSmall qsm1 qsm2) }) refl
where
go : {a x y quot quot2 : } (a *N quot2 +N zero y) (a *N quot +N zero x) a *N (quot +N quot2) +N zero x +N y
go {a} {x} {y} {quot} {quot2} pr1 pr2 rewrite Semiring.sumZeroRight Semiring (a *N quot) = identityOfIndiscernablesLeft _≡_ t (equalityCommutative (Semiring.sumZeroRight Semiring (a *N (quot +N quot2))))
where
t : a *N (quot +N quot2) x +N y
t rewrite Semiring.sumZeroRight Semiring (a *N quot2) = transitivity (Semiring.+DistributesOver* Semiring a quot quot2) p
where
s : a *N quot +N a *N quot2 x +N a *N quot2
s = applyEquality (λ n n +N a *N quot2) pr2
r : x +N a *N quot2 x +N y
r = applyEquality (λ n x +N n) pr1
p : a *N quot +N a *N quot2 x +N y
p = transitivity s r
quotSmall : ((0 <N a) || ((0 a) && (xDivA 0))) ((0 <N a) || ((0 a) && (quot 0))) (0 <N a) || ((0 a) && (xDivA +N quot 0))
quotSmall (inl x1) (inl x2) = inl x1
quotSmall (inl x1) (inr x2) = inl x1
quotSmall (inr x1) (inl x2) = inl x2
quotSmall (inr (a=0 ,, bl)) (inr (_ ,, bl2)) = inr (a=0 ,, ans)
where
ans : xDivA +N quot 0
ans rewrite bl | bl2 = refl
dividesBothImpliesDividesDifference : {a b c : } a b a c (c<b : c <N b) a (subtractionNResult.result (-N (inl c<b)))
dividesBothImpliesDividesDifference {zero} {b} {.0} prab (divides record { quot = quot ; rem = .0 ; pr = refl } refl) c<b = prab
dividesBothImpliesDividesDifference {succ a} {b} {c} (divides record { quot = bDivSA ; rem = .0 ; pr = pr } refl) (divides record { quot = cDivSA ; rem = .0 ; pr = pr2 } refl) c<b rewrite (Semiring.sumZeroRight Semiring (succ a *N cDivSA)) | (Semiring.sumZeroRight Semiring (succ a *N bDivSA)) = divides (record { quot = subtractionNResult.result bDivSA-cDivSA ; rem = 0 ; pr = identityOfIndiscernablesLeft _≡_ (identityOfIndiscernablesLeft _≡_ s (equalityCommutative q)) (equalityCommutative (Semiring.sumZeroRight Semiring _)) ; remIsSmall = inl (succIsPositive a) ; quotSmall = inl (succIsPositive a) }) refl
where
p : cDivSA <N bDivSA
p rewrite (equalityCommutative pr2) | (equalityCommutative pr) = cancelInequalityLeft {succ a} {cDivSA} {bDivSA} c<b
bDivSA-cDivSA : subtractionNResult cDivSA bDivSA (inl p)
bDivSA-cDivSA = -N {cDivSA} {bDivSA} (inl p)
la-ka = subtractionNResult.result (-N {succ a *N cDivSA} {succ a *N bDivSA} (inl (lessRespectsMultiplicationLeft cDivSA bDivSA (succ a) p (succIsPositive a))))
q : (succ a *N (subtractionNResult.result bDivSA-cDivSA)) la-ka
q = subtractProduct {succ a} {cDivSA} {bDivSA} (succIsPositive a) p
s : la-ka subtractionNResult.result (-N {c} {b} (inl c<b))
s = equivalentSubtraction (succ a *N cDivSA) b (succ a *N bDivSA) c (lessRespectsMultiplicationLeft cDivSA bDivSA (succ a) p (succIsPositive a)) c<b g
where
g : (succ a *N cDivSA) +N b (succ a *N bDivSA) +N c
g rewrite equalityCommutative pr2 | equalityCommutative pr = Semiring.commutative Semiring (cDivSA +N a *N cDivSA) (bDivSA +N a *N bDivSA)
euclidLemma1 : {a b : } (a<b : a <N b) (t : ) a +N b <N t a +N (subtractionNResult.result (-N (inl a<b))) <N t
euclidLemma1 {zero} {b} zero<b t b<t = b<t
euclidLemma1 {succ a} {b} sa<b t sa+b<t = identityOfIndiscernablesLeft _<N_ q (Semiring.commutative Semiring (subtractionNResult.result (-N (inl sa<b))) (succ a))
where
p : b <N t
p = TotalOrder.<Transitive TotalOrder (le a refl) sa+b<t
q : (subtractionNResult.result (-N (inl sa<b))) +N succ a <N t
q = identityOfIndiscernablesLeft _<N_ p (equalityCommutative (addMinus (inl sa<b)))
euclidLemma2 : {a b max : } (succ (a +N b) <N max) b <N max
euclidLemma2 {a} {b} {max} pr = lessTransitive {b} {succ (a +N b)} {max} (lemma a b) pr
where
lemma : (a b : ) b <N succ (a +N b)
lemma a b rewrite Semiring.commutative Semiring (succ a) b = addingIncreases b a
euclidLemma3 : {a b max : } (succ (succ (a +N b)) <N max) succ b <N max
euclidLemma3 {a} {b} {max} pr = euclidLemma2 {a} {succ b} {max} (identityOfIndiscernablesLeft _<N_ pr (applyEquality succ (equalityCommutative (succExtracts a b))))
euclidLemma4 : (a b c d h : ) (sa<b : (succ a) <N b) (pr : subtractionNResult.result (-N (inl sa<b)) *N c (succ a) *N d +N h) b *N c (succ a) *N (d +N c) +N h
euclidLemma4 a b zero d h sa<b pr rewrite Semiring.sumZeroRight Semiring d | Semiring.productZeroRight Semiring b | Semiring.productZeroRight Semiring (subtractionNResult.result (-N (inl sa<b))) = pr
euclidLemma4 a b (succ c) d h sa<b pr rewrite subtractProduct' (succIsPositive c) sa<b = transitivity q' r'
where
q : (succ c) *N b succ (a +N c *N succ a) +N ((succ a) *N d +N h)
q = moveOneSubtraction {succ (a +N c *N succ a)} {b +N c *N b} {(succ a) *N d +N h} {inl _} pr
q' : b *N succ c succ (a +N c *N succ a) +N ((succ a) *N d +N h)
q' rewrite multiplicationNIsCommutative b (succ c) = q
r' : ((succ c) *N succ a) +N (((succ a) *N d) +N h) ((succ a) *N (d +N succ c)) +N h
r' rewrite Semiring.+Associative Semiring ((succ c) *N succ a) ((succ a) *N d) h = applyEquality (λ t t +N h) {((succ c) *N succ a) +N ((succ a) *N d)} {(succ a) *N (d +N succ c)} (go (succ c) (succ a) d)
where
go' : (a b c : ) b *N a +N b *N c b *N (c +N a)
go : (a b c : ) a *N b +N b *N c b *N (c +N a)
go a b c rewrite multiplicationNIsCommutative a b = go' a b c
go' a b c rewrite Semiring.commutative Semiring (b *N a) (b *N c) = equalityCommutative (Semiring.+DistributesOver* Semiring b c a)
euclidLemma5 : (a b c d h : ) (sa<b : (succ a) <N b) (pr : subtractionNResult.result (-N (inl sa<b)) *N c +N h (succ a) *N d) (succ a) *N (d +N c) b *N c +N h
euclidLemma5 a b c d h sa<b pr with (-N (inl sa<b))
euclidLemma5 a b zero d h sa<b pr | record { result = result ; pr = sub } rewrite Semiring.sumZeroRight Semiring d | Semiring.productZeroRight Semiring b | Semiring.productZeroRight Semiring result = equalityCommutative pr
euclidLemma5 a b (succ c) d h sa<b pr | record { result = result ; pr = sub } rewrite subtractProduct' (succIsPositive c) sa<b | equalityCommutative sub = pv''
where
p : succ a *N d result *N succ c +N h
p = equalityCommutative pr
p' : a *N succ c +N succ a *N d (a *N succ c) +N ((result *N succ c) +N h)
p' = applyEquality (λ t a *N succ c +N t) p
p'' : a *N succ c +N succ a *N d (a *N succ c +N result *N succ c) +N h
p'' rewrite equalityCommutative (Semiring.+Associative Semiring (a *N succ c) (result *N succ c) h) = p'
p''' : a *N succ c +N succ a *N d (a +N result) *N succ c +N h
p''' rewrite multiplicationNIsCommutative (a +N result) (succ c) | Semiring.+DistributesOver* Semiring (succ c) a result | multiplicationNIsCommutative (succ c) a | multiplicationNIsCommutative (succ c) result = p''
pv : c +N (a *N succ c +N succ a *N d) (c +N (a +N result) *N succ c) +N h
pv rewrite equalityCommutative (Semiring.+Associative Semiring c ((a +N result) *N succ c) h) = applyEquality (λ t c +N t) p'''
pv' : (succ c) +N (a *N succ c +N succ a *N d) succ ((c +N (a +N result) *N succ c) +N h)
pv' = applyEquality succ pv
pv'' : (succ a) *N (d +N succ c) succ ((c +N (a +N result) *N succ c) +N h)
pv'' = identityOfIndiscernablesLeft _≡_ pv' (go a c d)
where
go : (a c d : ) (succ c) +N (a *N succ c +N ((succ a) *N d)) (succ a) *N (d +N succ c)
go a c d rewrite Semiring.+Associative Semiring (succ c) (a *N succ c) ((succ a) *N d) = go'
where
go' : (succ a) *N (succ c) +N (succ a) *N d (succ a) *N (d +N succ c)
go' rewrite Semiring.commutative Semiring d (succ c) = equalityCommutative (Semiring.+DistributesOver* Semiring (succ a) (succ c) d)
euclid : (a b : ) extendedHcf a b
euclid a b = inducted (succ a +N b) a b (a<SuccA (a +N b))
where
P : Set
P sum = (a b : ) a +N b <N sum extendedHcf a b
go'' : {a b : } (maxsum : ) (a <N b) (a +N b <N maxsum) ( y y <N maxsum P y) extendedHcf a b
go'' {zero} {b} maxSum zero<b b<maxsum indHyp = hcfZero b
go'' {1} {b} maxSum 1<b b<maxsum indHyp = hcfOne b
go'' {succ (succ a)} {b} maxSum ssa<b ssa+b<maxsum indHyp with (indHyp (succ b) (euclidLemma3 {a} {b} {maxSum} ssa+b<maxsum)) (subtractionNResult.result (-N (inl ssa<b))) (succ (succ a)) (identityOfIndiscernablesLeft _<N_ (a<SuccA b) (equalityCommutative (addMinus (inl ssa<b))))
go'' {succ (succ a)} {b} maxSum ssa<b ssa+b<maxsum indHyp | record { hcf = record { c = c ; c|a = c|a ; c|b = c|b ; hcf = hcf } ; extended1 = extended1 ; extended2 = extended2 ; extendedProof = inl extendedProof } = record { hcf = record { c = c ; c|a = c|b ; c|b = hcfDivB'' ; hcf = λ div prDivSSA prDivB hcf div (dividesBothImpliesDividesDifference prDivB prDivSSA ssa<b) prDivSSA } ; extended2 = extended1; extended1 = extended2 +N extended1 ; extendedProof = inr (equalityCommutative (euclidLemma4 (succ a) b extended1 extended2 c ssa<b extendedProof)) }
where
hcfDivB : c ((succ (succ a)) +N (subtractionNResult.result (-N (inl ssa<b))))
hcfDivB = dividesBothImpliesDividesSum {c} {succ (succ a)} { subtractionNResult.result (-N (inl ssa<b))} c|b c|a
hcfDivB' : c ((subtractionNResult.result (-N (inl ssa<b))) +N (succ (succ a)))
hcfDivB' = identityOfIndiscernablesRight __ hcfDivB (Semiring.commutative Semiring (succ (succ a)) ( subtractionNResult.result (-N (inl ssa<b))))
hcfDivB'' : c b
hcfDivB'' = identityOfIndiscernablesRight __ hcfDivB' (addMinus (inl ssa<b))
go'' {succ (succ a)} {b} maxSum ssa<b ssa+b<maxsum indHyp | record { hcf = record { c = c ; c|a = c|a ; c|b = c|b ; hcf = hcf } ; extended1 = extended1 ; extended2 = extended2 ; extendedProof = inr extendedProof } = record { hcf = record { c = c ; c|a = c|b ; c|b = hcfDivB'' ; hcf = λ div prDivSSA prDivB hcf div (dividesBothImpliesDividesDifference prDivB prDivSSA ssa<b) prDivSSA } ; extended2 = extended1; extended1 = extended2 +N extended1 ; extendedProof = inl (euclidLemma5 (succ a) b extended1 extended2 c ssa<b extendedProof) }
where
hcfDivB : c ((succ (succ a)) +N (subtractionNResult.result (-N (inl ssa<b))))
hcfDivB = dividesBothImpliesDividesSum {c} {succ (succ a)} { subtractionNResult.result (-N (inl ssa<b))} c|b c|a
hcfDivB' : c ((subtractionNResult.result (-N (inl ssa<b))) +N (succ (succ a)))
hcfDivB' = identityOfIndiscernablesRight __ hcfDivB (Semiring.commutative Semiring (succ (succ a)) (subtractionNResult.result (-N (inl ssa<b))))
hcfDivB'' : c b
hcfDivB'' = identityOfIndiscernablesRight __ hcfDivB' (addMinus (inl ssa<b))
go' : (maxSum a b : ) (a +N b <N maxSum) ( y y <N maxSum P y) extendedHcf a b
go' maxSum a b a+b<maxsum indHyp with totality a b
go' maxSum a b a+b<maxsum indHyp | inl (inl a<b) = go'' maxSum a<b a+b<maxsum indHyp
go' maxSum a b a+b<maxsum indHyp | inl (inr b<a) = reverseHCF (go'' maxSum b<a (identityOfIndiscernablesLeft _<N_ a+b<maxsum (Semiring.commutative Semiring a b)) indHyp)
go' maxSum a .a _ indHyp | inr refl = record { hcf = record { c = a ; c|a = aDivA a ; c|b = aDivA a ; hcf = λ _ _ z z } ; extended1 = 0 ; extended2 = 1 ; extendedProof = inr s}
where
s : a *N zero +N a a *N 1
s rewrite multiplicationNIsCommutative a zero | Semiring.productOneRight Semiring a = refl
go : x ( y y <N x P y) P x
go maxSum indHyp = λ a b a+b<maxSum go' maxSum a b a+b<maxSum indHyp
inducted : x P x
inducted = rec <NWellfounded P go

View File

@@ -105,7 +105,7 @@ lessImpliesNotEqual {a} {.a} prAB refl = TotalOrder.irreflexive TotalOrder pr
-NIsDecreasing {a} {b} prAB with (-N (inl prAB))
-NIsDecreasing {a} {b} (le x proof) | record { result = result ; pr = pr } = record { x = a ; proof = pr }
equalityN : (a b : ) Sg Bool (λ truth if truth then a b else Unit)
equalityN : (a b : ) Sg Bool (λ truth if truth then a b else True)
equalityN zero zero = ( BoolTrue , refl )
equalityN zero (succ b) = ( BoolFalse , record {} )
equalityN (succ a) zero = ( BoolFalse , record {} )

View File

@@ -9,6 +9,7 @@ open import Numbers.Primes.PrimeNumbers
open import WellFoundedInduction
open import Semirings.Definition
open import Orders
open import Numbers.Naturals.EuclideanAlgorithm
module Numbers.Primes.IntegerFactorisation where
open TotalOrder TotalOrder

View File

@@ -14,94 +14,13 @@ open import Vectors
open import Maybe
open import WithK
open import Semirings.Definition
open import Numbers.Naturals.EuclideanAlgorithm
module Numbers.Primes.PrimeNumbers where
open TotalOrder TotalOrder
open Semiring Semiring
record divisionAlgResult (a : ) (b : ) : Set where
field
quot :
rem :
pr : a *N quot +N rem b
remIsSmall : (rem <N a) || (a 0)
quotSmall : (0 <N a) || ((0 a) && (quot 0))
divAlgLessLemma : (a b : ) (0 <N a) (r : divisionAlgResult a b) (divisionAlgResult.quot r 0) || (divisionAlgResult.rem r <N b)
divAlgLessLemma zero b pr _ = exFalso (TotalOrder.irreflexive TotalOrder pr)
divAlgLessLemma (succ a) b _ record { quot = zero ; rem = a%b ; pr = pr ; remIsSmall = remIsSmall } = inl refl
divAlgLessLemma (succ a) b _ record { quot = (succ a/b) ; rem = a%b ; pr = pr ; remIsSmall = remIsSmall } = inr record { x = a/b +N a *N succ (a/b) ; proof = pr }
modUniqueLemma : {rem1 rem2 a : } (quot1 quot2 : ) rem1 <N a rem2 <N a a *N quot1 +N rem1 a *N quot2 +N rem2 rem1 rem2
modUniqueLemma {rem1} {rem2} {a} zero zero rem1<a rem2<a pr rewrite Semiring.productZeroRight Semiring a = pr
modUniqueLemma {rem1} {rem2} {a} zero (succ quot2) rem1<a rem2<a pr rewrite Semiring.productZeroRight Semiring a | pr | multiplicationNIsCommutative a (succ quot2) | equalityCommutative (Semiring.+Associative Semiring a (quot2 *N a) rem2) = exFalso (cannotAddAndEnlarge' {a} {quot2 *N a +N rem2} rem1<a)
modUniqueLemma {rem1} {rem2} {a} (succ quot1) zero rem1<a rem2<a pr rewrite Semiring.productZeroRight Semiring a | equalityCommutative pr | multiplicationNIsCommutative a (succ quot1) | equalityCommutative (Semiring.+Associative Semiring a (quot1 *N a) rem1) = exFalso (cannotAddAndEnlarge' {a} {quot1 *N a +N rem1} rem2<a)
modUniqueLemma {rem1} {rem2} {a} (succ quot1) (succ quot2) rem1<a rem2<a pr rewrite multiplicationNIsCommutative a (succ quot1) | multiplicationNIsCommutative a (succ quot2) | equalityCommutative (Semiring.+Associative Semiring a (quot1 *N a) rem1) | equalityCommutative (Semiring.+Associative Semiring a (quot2 *N a) rem2) = modUniqueLemma {rem1} {rem2} {a} quot1 quot2 rem1<a rem2<a (go {a}{quot1}{rem1}{quot2}{rem2} pr)
where
go : {a quot1 rem1 quot2 rem2 : } (a +N (quot1 *N a +N rem1) a +N (quot2 *N a +N rem2)) a *N quot1 +N rem1 a *N quot2 +N rem2
go {a} {quot1} {rem1} {quot2} {rem2} pr rewrite multiplicationNIsCommutative quot1 a | multiplicationNIsCommutative quot2 a = canSubtractFromEqualityLeft {a} pr
modIsUnique : {a b : } (div1 div2 : divisionAlgResult a b) divisionAlgResult.rem div1 divisionAlgResult.rem div2
modIsUnique {zero} {b} record { quot = quot1 ; rem = rem1 ; pr = pr1 ; remIsSmall = remIsSmall1 } record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = remIsSmall } = transitivity pr1 (equalityCommutative pr)
modIsUnique {succ a} {b} record { quot = quot1 ; rem = rem1 ; pr = pr1 ; remIsSmall = (inl y) } record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = (inl x) } = modUniqueLemma {rem1} {rem} {succ a} quot1 quot y x (transitivity pr1 (equalityCommutative pr))
modIsUnique {succ a} {b} record { quot = quot1 ; rem = rem1 ; pr = pr1 ; remIsSmall = (inr ()) } record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = (inl x) }
modIsUnique {succ a} {b} record { quot = quot1 ; rem = rem1 ; pr = pr1 ; remIsSmall = remIsSmall1 } record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = (inr ()) }
transferAddition : (a b c : ) (a +N b) +N c (a +N c) +N b
transferAddition a b c rewrite equalityCommutative (Semiring.+Associative Semiring a b c) = p a b c
where
p : (a b c : ) a +N (b +N c) (a +N c) +N b
p a b c = transitivity (applyEquality (a +N_) (Semiring.commutative Semiring b c)) (Semiring.+Associative Semiring a c b)
divisionAlgLemma : (x b : ) x *N zero +N b b
divisionAlgLemma x b rewrite (Semiring.productZeroRight Semiring x) = refl
divisionAlgLemma2 : (x b : ) (x b) x *N succ zero +N zero b
divisionAlgLemma2 x b pr rewrite (Semiring.productOneRight Semiring x) = equalityCommutative (transitivity (equalityCommutative pr) (equalityCommutative (Semiring.sumZeroRight Semiring x)))
divisionAlgLemma3 : {a x : } (p : succ a <N succ x) (subtractionNResult.result (-N (inl p))) <N (succ x)
divisionAlgLemma3 {a} {x} p = -NIsDecreasing {a} {succ x} p
divisionAlgLemma4 : (p a q : ) ((p +N a *N p) +N q) +N succ a succ ((p +N a *N succ p) +N q)
divisionAlgLemma4 p a q = ans
where
r : ((p +N a *N p) +N q) +N succ a succ (((p +N a *N p) +N q) +N a)
ans : ((p +N a *N p) +N q) +N succ a succ ((p +N a *N succ p) +N q)
s : ((p +N a *N p) +N q) +N a (p +N a *N succ p) +N q
t : (p +N a *N p) +N a p +N a *N succ p
s = transitivity (transferAddition (p +N a *N p) q a) (applyEquality (λ i i +N q) t)
ans = identityOfIndiscernablesRight _≡_ r (applyEquality succ s)
r = succExtracts ((p +N a *N p) +N q) a
t = transitivity (equalityCommutative (Semiring.+Associative Semiring p (a *N p) a)) (applyEquality (λ n p +N n) (equalityCommutative (transitivity (multiplicationNIsCommutative a (succ p)) (transitivity (Semiring.commutative Semiring a _) (applyEquality (_+N a) (multiplicationNIsCommutative p _))))))
divisionAlg : (a : ) (b : ) divisionAlgResult a b
divisionAlg zero = λ b record { quot = zero ; rem = b ; pr = refl ; remIsSmall = inr refl ; quotSmall = inr (record { fst = refl ; snd = refl }) }
divisionAlg (succ a) = rec <NWellfounded (λ n divisionAlgResult (succ a) n) go
where
go : (x : ) (indHyp : (y : ) (y<x : y <N x) divisionAlgResult (succ a) y)
divisionAlgResult (succ a) x
go zero prop = record { quot = zero ; rem = zero ; pr = divisionAlgLemma (succ a) zero ; remIsSmall = inl (succIsPositive a) ; quotSmall = inl (succIsPositive a) }
go (succ x) indHyp with totality (succ a) (succ x)
go (succ x) indHyp | inl (inl sa<sx) with indHyp (subtractionNResult.result (-N (inl sa<sx))) (divisionAlgLemma3 sa<sx)
... | record { quot = prevQuot ; rem = prevRem ; pr = prevPr ; remIsSmall = smallRem } = p
where
p : divisionAlgResult (succ a) (succ x)
addedA : (succ a *N prevQuot +N prevRem) +N (succ a) subtractionNResult.result (-N (inl sa<sx)) +N (succ a)
addedA' : (succ a *N prevQuot +N prevRem) +N succ a succ x
addedA'' : (succ a *N succ prevQuot) +N prevRem succ x
addedA''' : succ ((prevQuot +N a *N succ prevQuot) +N prevRem) succ x
addedA''' = identityOfIndiscernablesLeft _≡_ addedA'' refl
p = record { quot = succ prevQuot ; rem = prevRem ; pr = addedA''' ; remIsSmall = smallRem ; quotSmall = inl (succIsPositive a) }
addedA = applyEquality (λ n n +N succ a) prevPr
addedA' = identityOfIndiscernablesRight _≡_ addedA (addMinus {succ a} {succ x} (inl sa<sx))
addedA'' = identityOfIndiscernablesLeft _≡_ addedA' (divisionAlgLemma4 prevQuot a prevRem)
go (succ x) indHyp | inr (sa=sx) = record { quot = succ zero ; rem = zero ; pr = divisionAlgLemma2 (succ a) (succ x) sa=sx ; remIsSmall = inl (succIsPositive a) ; quotSmall = inl (succIsPositive a) }
go (succ x) indHyp | inl (inr (sx<sa)) = record { quot = zero ; rem = succ x ; pr = divisionAlgLemma (succ a) (succ x) ; remIsSmall = inl sx<sa ; quotSmall = inl (succIsPositive a) }
data __ : Set where
divides : {a b : } (res : divisionAlgResult a b) divisionAlgResult.rem res zero a b
dividesEqualityLemma'' : {a b : } (quot1 quot2 : ) (quot1 quot2) (rem : ) (pr1 : (quot1 +N a *N quot1) +N rem b) (pr2 : (quot2 +N a *N quot2) +N rem b) (y : rem <N succ a) (x1 : zero <N succ a) record { quot = quot1 ; rem = rem ; pr = pr1 ; remIsSmall = inl y ; quotSmall = inl x1 } record { quot = quot2 ; rem = rem ; pr = pr2 ; remIsSmall = inl y ; quotSmall = inl x1}
dividesEqualityLemma'' {a} {b} q1 .q1 refl rem pr1 pr2 y x1 rewrite reflRefl pr1 pr2 = refl
@@ -140,12 +59,6 @@ dividesEquality (divides res1 x1) (divides res2 x2) rewrite dividesEqualityPr re
data notDiv : Set where
doesNotDivide : {a b : } (res : divisionAlgResult a b) 0 <N divisionAlgResult.rem res notDiv a b
zeroDividesNothing : (a : ) zero succ a False
zeroDividesNothing a (divides record { quot = quot ; rem = rem ; pr = pr } x) = naughtE p
where
p : zero succ a
p = transitivity (equalityCommutative x) pr
twoDividesFour : succ (succ zero) succ (succ (succ (succ zero)))
twoDividesFour = divides {(succ (succ zero))} {succ (succ (succ (succ zero)))} (record { quot = succ (succ zero) ; rem = zero ; pr = refl ; remIsSmall = inl (succIsPositive 1) ; quotSmall = inl (succIsPositive 1) }) refl
@@ -208,13 +121,6 @@ compositeImpliesNotPrime (succ (succ m)) (succ (succ p)) _ mLessP mDivP pPrime =
fourIsNotPrime : Prime 4 False
fourIsNotPrime = compositeImpliesNotPrime (succ (succ zero)) (succ (succ (succ (succ zero)))) (le zero refl) (le (succ zero) refl) twoDividesFour
record hcfData (a b : ) : Set where
field
c :
c|a : c a
c|b : c b
hcf : x x a x b x c
record Coprime (a : ) (b : ) : Set where
field
hcf : hcfData a b
@@ -252,53 +158,6 @@ allNumbersLessThanDescending (succ n) = record { a = n ; a<n = le zero refl } ,-
allNumbersLessThan : (n : ) Vec (numberLessThan n) n
allNumbersLessThan n = vecRev (allNumbersLessThanDescending n)
positiveTimes : {a b : } (succ a *N succ b <N succ a) False
positiveTimes {a} {b} pr = zeroNeverGreater f'
where
g : succ a *N succ b <N succ a *N succ 0
g rewrite multiplicationNIsCommutative a 1 | Semiring.commutative Semiring a 0 = pr
f : succ b <N succ 0
f = cancelInequalityLeft {succ a} {succ b} g
f' : b <N 0
f' = canRemoveSuccFrom<N f
biggerThanCantDivideLemma : {a b : } (a <N b) (b a) a 0
biggerThanCantDivideLemma {zero} {b} a<b b|a = refl
biggerThanCantDivideLemma {succ a} {zero} a<b (divides record { quot = quot ; rem = .0 ; pr = pr ; remIsSmall = remIsSmall ; quotSmall = (inl (le x ())) } refl)
biggerThanCantDivideLemma {succ a} {zero} a<b (divides record { quot = quot ; rem = .0 ; pr = () ; remIsSmall = remIsSmall ; quotSmall = (inr x) } refl)
biggerThanCantDivideLemma {succ a} {succ b} a<b (divides record { quot = zero ; rem = .0 ; pr = pr ; remIsSmall = remIsSmall ; quotSmall = quotSmall } refl) rewrite Semiring.commutative Semiring (b *N zero) 0 | multiplicationNIsCommutative b 0 = exFalso (naughtE pr)
biggerThanCantDivideLemma {succ a} {succ b} a<b (divides record { quot = (succ quot) ; rem = .0 ; pr = pr ; remIsSmall = remIsSmall ; quotSmall = quotSmall } refl) rewrite Semiring.commutative Semiring (quot +N b *N succ quot) 0 | equalityCommutative pr = exFalso (positiveTimes {b} {quot} a<b)
biggerThanCantDivide : {a b : } (x : ) (TotalOrder.max TotalOrder a b) <N x (x a) (x b) (a 0) && (b 0)
biggerThanCantDivide {a} {b} x pr x|a x|b with totality a b
biggerThanCantDivide {a} {b} x pr x|a x|b | inl (inl a<b) = exFalso (zeroNeverGreater f')
where
f : b 0
f = biggerThanCantDivideLemma pr x|b
f' : a <N 0
f' rewrite equalityCommutative f = a<b
biggerThanCantDivide {a} {b} x pr x|a x|b | inl (inr b<a) = exFalso (zeroNeverGreater f')
where
f : a 0
f = biggerThanCantDivideLemma pr x|a
f' : b <N 0
f' rewrite equalityCommutative f = b<a
biggerThanCantDivide {a} {b} x pr x|a x|b | inr a=b = (transitivity a=b f ,, f)
where
f : b 0
f = biggerThanCantDivideLemma pr x|b
aDivA : (a : ) a a
aDivA zero = divides (record { quot = 0 ; rem = 0 ; pr = equalityCommutative (oneTimesPlusZero zero) ; remIsSmall = inr refl ; quotSmall = inr (refl ,, refl) }) refl
aDivA (succ a) = divides (record { quot = 1 ; rem = 0 ; pr = equalityCommutative (oneTimesPlusZero (succ a)) ; remIsSmall = inl (succIsPositive a) ; quotSmall = inl (succIsPositive a) }) refl
aDivZero : (a : ) a zero
aDivZero zero = aDivA zero
aDivZero (succ a) = divides (record { quot = zero ; rem = zero ; pr = lemma (succ a) ; remIsSmall = inl (succIsPositive a) ; quotSmall = inl (succIsPositive a) }) refl
where
lemma : (b : ) b *N zero +N zero zero
lemma b rewrite (Semiring.sumZeroRight Semiring (b *N zero)) = Semiring.productZeroRight Semiring b
maxDivides : (a b : ) ((TotalOrder.max TotalOrder a b) a) (TotalOrder.max TotalOrder a b) b (((a 0) && (0 <N b)) || ((b 0) && (0 <N a))) || (a b)
maxDivides a b max|a max|b with totality a b
maxDivides a b max|a max|b | inl (inl a<b) = inl (inl (record { fst = gg ; snd = identityOfIndiscernablesLeft _<N_ a<b gg}))
@@ -332,203 +191,6 @@ extensionalHCFEquality : {a b : } → {h1 h2 : extensionalHCF a b} → (exten
extensionalHCFEquality {a} {b} {record { c = c1 ; c|a = c|a1 ; c|b = c|b1 ; hcfExtension = hcfExtension1 }} {record { c = c2 ; c|a = c|a2 ; c|b = c|b2 ; hcfExtension = hcfExtension2 }} pr rewrite pr = {!!}
-}
record extendedHcf (a b : ) : Set where
field
hcf : hcfData a b
c :
c = hcfData.c hcf
field
extended1 :
extended2 :
extendedProof : (a *N extended1 b *N extended2 +N c) || (a *N extended1 +N c b *N extended2)
divEqualityLemma1 : {a b c : } b zero b *N c +N 0 a a b
divEqualityLemma1 {a} {.0} {c} refl pr = equalityCommutative pr
divEquality : {a b : } a b b a a b
divEquality {a} {b} (divides record { quot = quotAB ; rem = .0 ; pr = prAB ; remIsSmall = _ ; quotSmall = quotSmallAB } refl) (divides record { quot = quot ; rem = .0 ; pr = pr ; remIsSmall = _ ; quotSmall = (inl x) } refl) rewrite Semiring.commutative Semiring (b *N quot) 0 | Semiring.commutative Semiring (a *N quotAB) 0 | equalityCommutative pr | equalityCommutative (Semiring.*Associative Semiring b quot quotAB) = res
where
lem : {b r : } b *N r b (0 <N b) r 1
lem {zero} {r} pr ()
lem {succ b} {zero} pr _ rewrite multiplicationNIsCommutative b 0 = exFalso (naughtE pr)
lem {succ b} {succ zero} pr _ = refl
lem {succ b} {succ (succ r)} pr _ rewrite multiplicationNIsCommutative b (succ (succ r)) | Semiring.commutative Semiring (succ r) (b +N (b +N r *N b)) | equalityCommutative (Semiring.+Associative Semiring b (b +N r *N b) (succ r)) | Semiring.commutative Semiring (b +N r *N b) (succ r) = exFalso (cannotAddAndEnlarge'' {succ b} pr)
p : quot *N quotAB 1
p = lem prAB x
q : quot 1
q = _&&_.fst (productOneImpliesOperandsOne p)
res : b *N quot b
res rewrite q | multiplicationNIsCommutative b 1 | Semiring.commutative Semiring b 0 = refl
divEquality {.0} {.0} (divides record { quot = quotAB ; rem = .0 ; pr = prAB ; remIsSmall = _ ; quotSmall = quotSmallAB } refl) (divides record { quot = quot ; rem = .0 ; pr = refl ; remIsSmall = _ ; quotSmall = (inr (refl ,, snd)) } refl) = refl
hcfWelldefined : {a b : } (ab : hcfData a b) (ab' : hcfData a b) (hcfData.c ab hcfData.c ab')
hcfWelldefined {a} {b} record { c = c ; c|a = c|a ; c|b = c|b ; hcf = hcf } record { c = c' ; c|a = c|a' ; c|b = c|b' ; hcf = hcf' } with hcf c' c|a' c|b'
... | c'DivC with hcf' c c|a c|b
... | cDivC' = divEquality cDivC' c'DivC
reverseHCF : {a b : } (ab : extendedHcf a b) extendedHcf b a
reverseHCF {a} {b} record { hcf = record { c = c ; c|a = c|a ; c|b = c|b ; hcf = hcf } ; extended1 = extended1 ; extended2 = extended2 ; extendedProof = (inl x) } = record { hcf = record { c = c ; c|a = c|b ; c|b = c|a ; hcf = λ x z z₁ hcf x z₁ z } ; extended1 = extended2 ; extended2 = extended1 ; extendedProof = inr (equalityCommutative x) }
reverseHCF {a} {b} record { hcf = record { c = c ; c|a = c|a ; c|b = c|b ; hcf = hcf } ; extended1 = extended1 ; extended2 = extended2 ; extendedProof = (inr x) } = record { hcf = record { c = c ; c|a = c|b ; c|b = c|a ; hcf = λ x z z₁ hcf x z₁ z } ; extended1 = extended2 ; extended2 = extended1 ; extendedProof = inl (equalityCommutative x) }
oneDivN : (a : ) 1 a
oneDivN a = divides (record { quot = a ; rem = zero ; pr = pr ; remIsSmall = inl (succIsPositive zero) ; quotSmall = inl (le zero refl) }) refl
where
pr : (a +N zero) +N zero a
pr rewrite Semiring.sumZeroRight Semiring (a +N zero) = Semiring.sumZeroRight Semiring a
hcfZero : (a : ) extendedHcf zero a
hcfZero a = record { hcf = record { c = a ; c|a = aDivZero a ; c|b = aDivA a ; hcf = λ _ _ p p } ; extended1 = 0 ; extended2 = 1 ; extendedProof = inr (equalityCommutative (Semiring.productOneRight Semiring a))}
hcfOne : (a : ) extendedHcf 1 a
hcfOne a = record { hcf = record { c = 1 ; c|a = aDivA 1 ; c|b = oneDivN a ; hcf = λ _ z _ z } ; extended1 = 1 ; extended2 = 0 ; extendedProof = inl g }
where
g : 1 a *N 0 +N 1
g rewrite multiplicationNIsCommutative a 0 = refl
zeroIsValidRem : (a : ) (0 <N a) || (a 0)
zeroIsValidRem zero = inr refl
zeroIsValidRem (succ a) = inl (succIsPositive a)
dividesBothImpliesDividesSum : {a x y : } a x a y a (x +N y)
dividesBothImpliesDividesSum {a} {x} {y} (divides record { quot = xDivA ; rem = .0 ; pr = prA ; quotSmall = qsm1 } refl) (divides record { quot = quot ; rem = .0 ; pr = pr ; quotSmall = qsm2 } refl) = divides (record { quot = xDivA +N quot ; rem = 0 ; pr = go {a} {x} {y} {xDivA} {quot} pr prA ; remIsSmall = zeroIsValidRem a ; quotSmall = (quotSmall qsm1 qsm2) }) refl
where
go : {a x y quot quot2 : } (a *N quot2 +N zero y) (a *N quot +N zero x) a *N (quot +N quot2) +N zero x +N y
go {a} {x} {y} {quot} {quot2} pr1 pr2 rewrite Semiring.sumZeroRight Semiring (a *N quot) = identityOfIndiscernablesLeft _≡_ t (equalityCommutative (Semiring.sumZeroRight Semiring (a *N (quot +N quot2))))
where
t : a *N (quot +N quot2) x +N y
t rewrite Semiring.sumZeroRight Semiring (a *N quot2) = transitivity (Semiring.+DistributesOver* Semiring a quot quot2) p
where
s : a *N quot +N a *N quot2 x +N a *N quot2
s = applyEquality (λ n n +N a *N quot2) pr2
r : x +N a *N quot2 x +N y
r = applyEquality (λ n x +N n) pr1
p : a *N quot +N a *N quot2 x +N y
p = transitivity s r
quotSmall : ((0 <N a) || ((0 a) && (xDivA 0))) ((0 <N a) || ((0 a) && (quot 0))) (0 <N a) || ((0 a) && (xDivA +N quot 0))
quotSmall (inl x1) (inl x2) = inl x1
quotSmall (inl x1) (inr x2) = inl x1
quotSmall (inr x1) (inl x2) = inl x2
quotSmall (inr (a=0 ,, bl)) (inr (_ ,, bl2)) = inr (a=0 ,, ans)
where
ans : xDivA +N quot 0
ans rewrite bl | bl2 = refl
dividesBothImpliesDividesDifference : {a b c : } a b a c (c<b : c <N b) a (subtractionNResult.result (-N (inl c<b)))
dividesBothImpliesDividesDifference {zero} {b} {.0} prab (divides record { quot = quot ; rem = .0 ; pr = refl } refl) c<b = prab
dividesBothImpliesDividesDifference {succ a} {b} {c} (divides record { quot = bDivSA ; rem = .0 ; pr = pr } refl) (divides record { quot = cDivSA ; rem = .0 ; pr = pr2 } refl) c<b rewrite (Semiring.sumZeroRight Semiring (succ a *N cDivSA)) | (Semiring.sumZeroRight Semiring (succ a *N bDivSA)) = divides (record { quot = subtractionNResult.result bDivSA-cDivSA ; rem = 0 ; pr = identityOfIndiscernablesLeft _≡_ (identityOfIndiscernablesLeft _≡_ s (equalityCommutative q)) (equalityCommutative (Semiring.sumZeroRight Semiring _)) ; remIsSmall = inl (succIsPositive a) ; quotSmall = inl (succIsPositive a) }) refl
where
p : cDivSA <N bDivSA
p rewrite (equalityCommutative pr2) | (equalityCommutative pr) = cancelInequalityLeft {succ a} {cDivSA} {bDivSA} c<b
bDivSA-cDivSA : subtractionNResult cDivSA bDivSA (inl p)
bDivSA-cDivSA = -N {cDivSA} {bDivSA} (inl p)
la-ka = subtractionNResult.result (-N {succ a *N cDivSA} {succ a *N bDivSA} (inl (lessRespectsMultiplicationLeft cDivSA bDivSA (succ a) p (succIsPositive a))))
q : (succ a *N (subtractionNResult.result bDivSA-cDivSA)) la-ka
q = subtractProduct {succ a} {cDivSA} {bDivSA} (succIsPositive a) p
s : la-ka subtractionNResult.result (-N {c} {b} (inl c<b))
s = equivalentSubtraction (succ a *N cDivSA) b (succ a *N bDivSA) c (lessRespectsMultiplicationLeft cDivSA bDivSA (succ a) p (succIsPositive a)) c<b g
where
g : (succ a *N cDivSA) +N b (succ a *N bDivSA) +N c
g rewrite equalityCommutative pr2 | equalityCommutative pr = Semiring.commutative Semiring (cDivSA +N a *N cDivSA) (bDivSA +N a *N bDivSA)
euclidLemma1 : {a b : } (a<b : a <N b) (t : ) a +N b <N t a +N (subtractionNResult.result (-N (inl a<b))) <N t
euclidLemma1 {zero} {b} zero<b t b<t = b<t
euclidLemma1 {succ a} {b} sa<b t sa+b<t = identityOfIndiscernablesLeft _<N_ q (Semiring.commutative Semiring (subtractionNResult.result (-N (inl sa<b))) (succ a))
where
p : b <N t
p = TotalOrder.<Transitive TotalOrder (le a refl) sa+b<t
q : (subtractionNResult.result (-N (inl sa<b))) +N succ a <N t
q = identityOfIndiscernablesLeft _<N_ p (equalityCommutative (addMinus (inl sa<b)))
euclidLemma2 : {a b max : } (succ (a +N b) <N max) b <N max
euclidLemma2 {a} {b} {max} pr = lessTransitive {b} {succ (a +N b)} {max} (lemma a b) pr
where
lemma : (a b : ) b <N succ (a +N b)
lemma a b rewrite Semiring.commutative Semiring (succ a) b = addingIncreases b a
euclidLemma3 : {a b max : } (succ (succ (a +N b)) <N max) succ b <N max
euclidLemma3 {a} {b} {max} pr = euclidLemma2 {a} {succ b} {max} (identityOfIndiscernablesLeft _<N_ pr (applyEquality succ (equalityCommutative (succExtracts a b))))
euclidLemma4 : (a b c d h : ) (sa<b : (succ a) <N b) (pr : subtractionNResult.result (-N (inl sa<b)) *N c (succ a) *N d +N h) b *N c (succ a) *N (d +N c) +N h
euclidLemma4 a b zero d h sa<b pr rewrite Semiring.sumZeroRight Semiring d | Semiring.productZeroRight Semiring b | Semiring.productZeroRight Semiring (subtractionNResult.result (-N (inl sa<b))) = pr
euclidLemma4 a b (succ c) d h sa<b pr rewrite subtractProduct' (succIsPositive c) sa<b = transitivity q' r'
where
q : (succ c) *N b succ (a +N c *N succ a) +N ((succ a) *N d +N h)
q = moveOneSubtraction {succ (a +N c *N succ a)} {b +N c *N b} {(succ a) *N d +N h} {inl _} pr
q' : b *N succ c succ (a +N c *N succ a) +N ((succ a) *N d +N h)
q' rewrite multiplicationNIsCommutative b (succ c) = q
r' : ((succ c) *N succ a) +N (((succ a) *N d) +N h) ((succ a) *N (d +N succ c)) +N h
r' rewrite Semiring.+Associative Semiring ((succ c) *N succ a) ((succ a) *N d) h = applyEquality (λ t t +N h) {((succ c) *N succ a) +N ((succ a) *N d)} {(succ a) *N (d +N succ c)} (go (succ c) (succ a) d)
where
go' : (a b c : ) b *N a +N b *N c b *N (c +N a)
go : (a b c : ) a *N b +N b *N c b *N (c +N a)
go a b c rewrite multiplicationNIsCommutative a b = go' a b c
go' a b c rewrite Semiring.commutative Semiring (b *N a) (b *N c) = equalityCommutative (Semiring.+DistributesOver* Semiring b c a)
euclidLemma5 : (a b c d h : ) (sa<b : (succ a) <N b) (pr : subtractionNResult.result (-N (inl sa<b)) *N c +N h (succ a) *N d) (succ a) *N (d +N c) b *N c +N h
euclidLemma5 a b c d h sa<b pr with (-N (inl sa<b))
euclidLemma5 a b zero d h sa<b pr | record { result = result ; pr = sub } rewrite Semiring.sumZeroRight Semiring d | Semiring.productZeroRight Semiring b | Semiring.productZeroRight Semiring result = equalityCommutative pr
euclidLemma5 a b (succ c) d h sa<b pr | record { result = result ; pr = sub } rewrite subtractProduct' (succIsPositive c) sa<b | equalityCommutative sub = pv''
where
p : succ a *N d result *N succ c +N h
p = equalityCommutative pr
p' : a *N succ c +N succ a *N d (a *N succ c) +N ((result *N succ c) +N h)
p' = applyEquality (λ t a *N succ c +N t) p
p'' : a *N succ c +N succ a *N d (a *N succ c +N result *N succ c) +N h
p'' rewrite equalityCommutative (Semiring.+Associative Semiring (a *N succ c) (result *N succ c) h) = p'
p''' : a *N succ c +N succ a *N d (a +N result) *N succ c +N h
p''' rewrite multiplicationNIsCommutative (a +N result) (succ c) | Semiring.+DistributesOver* Semiring (succ c) a result | multiplicationNIsCommutative (succ c) a | multiplicationNIsCommutative (succ c) result = p''
pv : c +N (a *N succ c +N succ a *N d) (c +N (a +N result) *N succ c) +N h
pv rewrite equalityCommutative (Semiring.+Associative Semiring c ((a +N result) *N succ c) h) = applyEquality (λ t c +N t) p'''
pv' : (succ c) +N (a *N succ c +N succ a *N d) succ ((c +N (a +N result) *N succ c) +N h)
pv' = applyEquality succ pv
pv'' : (succ a) *N (d +N succ c) succ ((c +N (a +N result) *N succ c) +N h)
pv'' = identityOfIndiscernablesLeft _≡_ pv' (go a c d)
where
go : (a c d : ) (succ c) +N (a *N succ c +N ((succ a) *N d)) (succ a) *N (d +N succ c)
go a c d rewrite Semiring.+Associative Semiring (succ c) (a *N succ c) ((succ a) *N d) = go'
where
go' : (succ a) *N (succ c) +N (succ a) *N d (succ a) *N (d +N succ c)
go' rewrite Semiring.commutative Semiring d (succ c) = equalityCommutative (Semiring.+DistributesOver* Semiring (succ a) (succ c) d)
euclid : (a b : ) extendedHcf a b
euclid a b = inducted (succ a +N b) a b (a<SuccA (a +N b))
where
P : Set
P sum = (a b : ) a +N b <N sum extendedHcf a b
go'' : {a b : } (maxsum : ) (a <N b) (a +N b <N maxsum) ( y y <N maxsum P y) extendedHcf a b
go'' {zero} {b} maxSum zero<b b<maxsum indHyp = hcfZero b
go'' {1} {b} maxSum 1<b b<maxsum indHyp = hcfOne b
go'' {succ (succ a)} {b} maxSum ssa<b ssa+b<maxsum indHyp with (indHyp (succ b) (euclidLemma3 {a} {b} {maxSum} ssa+b<maxsum)) (subtractionNResult.result (-N (inl ssa<b))) (succ (succ a)) (identityOfIndiscernablesLeft _<N_ (a<SuccA b) (equalityCommutative (addMinus (inl ssa<b))))
go'' {succ (succ a)} {b} maxSum ssa<b ssa+b<maxsum indHyp | record { hcf = record { c = c ; c|a = c|a ; c|b = c|b ; hcf = hcf } ; extended1 = extended1 ; extended2 = extended2 ; extendedProof = inl extendedProof } = record { hcf = record { c = c ; c|a = c|b ; c|b = hcfDivB'' ; hcf = λ div prDivSSA prDivB hcf div (dividesBothImpliesDividesDifference prDivB prDivSSA ssa<b) prDivSSA } ; extended2 = extended1; extended1 = extended2 +N extended1 ; extendedProof = inr (equalityCommutative (euclidLemma4 (succ a) b extended1 extended2 c ssa<b extendedProof)) }
where
hcfDivB : c ((succ (succ a)) +N (subtractionNResult.result (-N (inl ssa<b))))
hcfDivB = dividesBothImpliesDividesSum {c} {succ (succ a)} { subtractionNResult.result (-N (inl ssa<b))} c|b c|a
hcfDivB' : c ((subtractionNResult.result (-N (inl ssa<b))) +N (succ (succ a)))
hcfDivB' = identityOfIndiscernablesRight __ hcfDivB (Semiring.commutative Semiring (succ (succ a)) ( subtractionNResult.result (-N (inl ssa<b))))
hcfDivB'' : c b
hcfDivB'' = identityOfIndiscernablesRight __ hcfDivB' (addMinus (inl ssa<b))
go'' {succ (succ a)} {b} maxSum ssa<b ssa+b<maxsum indHyp | record { hcf = record { c = c ; c|a = c|a ; c|b = c|b ; hcf = hcf } ; extended1 = extended1 ; extended2 = extended2 ; extendedProof = inr extendedProof } = record { hcf = record { c = c ; c|a = c|b ; c|b = hcfDivB'' ; hcf = λ div prDivSSA prDivB hcf div (dividesBothImpliesDividesDifference prDivB prDivSSA ssa<b) prDivSSA } ; extended2 = extended1; extended1 = extended2 +N extended1 ; extendedProof = inl (euclidLemma5 (succ a) b extended1 extended2 c ssa<b extendedProof) }
where
hcfDivB : c ((succ (succ a)) +N (subtractionNResult.result (-N (inl ssa<b))))
hcfDivB = dividesBothImpliesDividesSum {c} {succ (succ a)} { subtractionNResult.result (-N (inl ssa<b))} c|b c|a
hcfDivB' : c ((subtractionNResult.result (-N (inl ssa<b))) +N (succ (succ a)))
hcfDivB' = identityOfIndiscernablesRight __ hcfDivB (Semiring.commutative Semiring (succ (succ a)) (subtractionNResult.result (-N (inl ssa<b))))
hcfDivB'' : c b
hcfDivB'' = identityOfIndiscernablesRight __ hcfDivB' (addMinus (inl ssa<b))
go' : (maxSum a b : ) (a +N b <N maxSum) ( y y <N maxSum P y) extendedHcf a b
go' maxSum a b a+b<maxsum indHyp with totality a b
go' maxSum a b a+b<maxsum indHyp | inl (inl a<b) = go'' maxSum a<b a+b<maxsum indHyp
go' maxSum a b a+b<maxsum indHyp | inl (inr b<a) = reverseHCF (go'' maxSum b<a (identityOfIndiscernablesLeft _<N_ a+b<maxsum (Semiring.commutative Semiring a b)) indHyp)
go' maxSum a .a _ indHyp | inr refl = record { hcf = record { c = a ; c|a = aDivA a ; c|b = aDivA a ; hcf = λ _ _ z z } ; extended1 = 0 ; extended2 = 1 ; extendedProof = inr s}
where
s : a *N zero +N a a *N 1
s rewrite multiplicationNIsCommutative a zero | Semiring.productOneRight Semiring a = refl
go : x ( y y <N x P y) P x
go maxSum indHyp = λ a b a+b<maxSum go' maxSum a b a+b<maxSum indHyp
inducted : x P x
inducted = rec <NWellfounded P go
divisorIsSmaller : {a b : } a succ b succ b <N a False
divisorIsSmaller {a} {b} (divides record { quot = zero ; rem = .0 ; pr = pr } refl) sb<a rewrite Semiring.sumZeroRight Semiring (a *N zero) = go
where

View File

@@ -0,0 +1,27 @@
{-# OPTIONS --safe --warning=error --without-K #-}
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
open import Setoids.Setoids
open import Functions
open import Sets.EquivalenceRelations
open import Rings.Definition
open import Rings.Homomorphisms.Definition
open import Groups.Homomorphisms.Lemmas
open import Rings.IntegralDomains.Definition
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
module Rings.Associates.Definition {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ _*_ : A A A} {R : Ring S _+_ _*_} (intDom : IntegralDomain R) where
open import Rings.Units.Definition R
open Setoid S
open Ring R
open Equivalence eq
Associates : Rel A
Associates x y = Sg A (λ z Unit z && (x (y * z)))

View File

@@ -0,0 +1,44 @@
{-# OPTIONS --safe --warning=error --without-K #-}
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
open import Setoids.Setoids
open import Functions
open import Sets.EquivalenceRelations
open import Rings.Definition
open import Rings.Homomorphisms.Definition
open import Groups.Homomorphisms.Lemmas
open import Rings.IntegralDomains.Definition
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
module Rings.Associates.Lemmas {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ _*_ : A A A} {R : Ring S _+_ _*_} (intDom : IntegralDomain R) where
open import Rings.Divisible.Definition R
open import Rings.IntegralDomains.Lemmas intDom
open import Rings.Associates.Definition intDom
open import Rings.Units.Definition R
open Setoid S
open Ring R
open Equivalence eq
associatesEquiv : Equivalence Associates
Equivalence.reflexive associatesEquiv {x} = 1R , ((1R , identIsIdent) ,, symmetric (transitive *Commutative identIsIdent))
Equivalence.symmetric associatesEquiv {x} {y} (a , ((invA , prInv) ,, x=ya)) = invA , ((a , transitive *Commutative prInv) ,, transitive (symmetric identIsIdent) (transitive (*WellDefined (symmetric prInv) reflexive) (transitive *Commutative (transitive *Associative (*WellDefined (symmetric x=ya) reflexive)))))
Equivalence.transitive associatesEquiv {x} {y} {z} (a , ((invA , prInvA) ,, x=ya)) (b , ((invB , prInvB) ,, y=zb)) = (a * b) , (((invA * invB) , transitive *Associative (transitive (*WellDefined (transitive *Commutative *Associative) reflexive) (transitive (symmetric *Associative) (transitive (*WellDefined (transitive *Commutative prInvA) prInvB) identIsIdent)))) ,, transitive x=ya (transitive (*WellDefined y=zb reflexive) (transitive (symmetric *Associative) (*WellDefined reflexive *Commutative))))
associateImpliesMutualDivision : {a b : A} Associates a b a b
associateImpliesMutualDivision {a} {b} (x , ((invX , prInvX) ,, a=bx)) = invX , transitive (transitive (*WellDefined a=bx reflexive) (transitive (transitive (symmetric *Associative) (*WellDefined reflexive prInvX)) *Commutative)) identIsIdent
mutualDivisionImpliesAssociate : {a b : A} (a b) (b a) ((a 0R) False) Associates a b
mutualDivisionImpliesAssociate {a} {b} (r , ar=b) (s , bs=a) a!=0 = s , ((r , cancelIntDom {a = a} (transitive (transitive (transitive (*WellDefined reflexive *Commutative) (transitive *Associative (*WellDefined ar=b reflexive))) bs=a) (transitive (symmetric identIsIdent) *Commutative)) a!=0) ,, symmetric bs=a)
mutualDivisionImpliesAssociate' : {a b : A} (a b) (b a) (a 0R) Associates a b
mutualDivisionImpliesAssociate' {a} {b} (r , ar=b) (s , bs=a) a=0 = 1R , ((1R , identIsIdent) ,, transitive a=0 (symmetric (transitive (*WellDefined b=0 reflexive) (transitive *Commutative timesZero))))
where
b=0 : b 0R
b=0 = transitive (symmetric ar=b) (transitive (transitive *Commutative (*WellDefined reflexive a=0)) (timesZero {r}))

View File

@@ -0,0 +1,28 @@
{-# OPTIONS --safe --warning=error --without-K #-}
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
open import Setoids.Setoids
open import Functions
open import Sets.EquivalenceRelations
open import Rings.Definition
open import Rings.Homomorphisms.Definition
open import Groups.Homomorphisms.Lemmas
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
module Rings.Divisible.Definition {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ _*_ : A A A} (R : Ring S _+_ _*_) where
open Setoid S
open Equivalence eq
open Ring R
__ : Rel A
a b = Sg A (λ c (a * c) b)
divisibleWellDefined : {x y a b : A} (x y) (a b) x a y b
divisibleWellDefined x=y a=b (c , xc=a) = c , transitive (*WellDefined (symmetric x=y) reflexive) (transitive xc=a a=b)

View File

@@ -22,7 +22,7 @@ module Rings.EuclideanDomains.Definition {a b : _} {A : Set a} {S : Setoid {a} {
open Setoid S
open Ring R
record DivisionAlgorithmResult {norm : {a : A} ((a 0R) False) } {x y : A} (x!=0 : (x 0R) False) (y!=0 : (y 0R) False) : Set (a b) where
record DivisionAlgorithmResult (norm : {a : A} ((a 0R) False) ) {x y : A} (x!=0 : (x 0R) False) (y!=0 : (y 0R) False) : Set (a b) where
field
quotient : A
rem : A
@@ -34,4 +34,4 @@ record EuclideanDomain : Set (a ⊔ lsuc b) where
isIntegralDomain : IntegralDomain R
norm : {a : A} ((a 0R) False)
normSize : {a b : A} (a!=0 : (a 0R) False) (b!=0 : (b 0R) False) (c : A) b (a * c) (norm a!=0) ≤N (norm b!=0)
divisionAlg : {a b : A} (a!=0 : (a 0R) False) (b!=0 : (b 0R) False) DivisionAlgorithmResult {norm} a!=0 b!=0
divisionAlg : {a b : A} (a!=0 : (a 0R) False) (b!=0 : (b 0R) False) DivisionAlgorithmResult norm a!=0 b!=0

View File

@@ -17,6 +17,7 @@ open import Rings.IntegralDomains.Definition
open import Rings.IntegralDomains.Examples
open import Rings.EuclideanDomains.Definition
open import Fields.Fields
open import Fields.Lemmas
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)

View File

@@ -0,0 +1,30 @@
{-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae
open import Groups.Groups
open import Groups.Homomorphisms.Definition
open import Groups.Definition
open import Numbers.Naturals.Definition
open import Numbers.Naturals.Order
open import Setoids.Orders
open import Setoids.Setoids
open import Functions
open import Sets.EquivalenceRelations
open import Rings.Definition
open import Rings.Homomorphisms.Definition
open import Groups.Homomorphisms.Lemmas
open import Rings.IntegralDomains.Definition
open import Rings.IntegralDomains.Examples
open import Rings.EuclideanDomains.Definition
open import Fields.Fields
open import WellFoundedInduction
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
module Rings.EuclideanDomains.Lemmas {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ _*_ : A A A} {R : Ring S _+_ _*_} (E : EuclideanDomain R) where
open import Rings.PrincipalIdealDomain R
open import Rings.Ideals.Principal.Definition R
euclideanDomainIsPid : {c : _} PrincipalIdealDomain {c}
euclideanDomainIsPid ideal = {!!}

View File

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

View File

@@ -25,6 +25,7 @@ open Equivalence eq
open Group additiveGroup
open import Rings.Lemmas R
open import Rings.Divisible.Definition R
record Ideal {c : _} (pred : A Set c) : Set (a b c) where
field
@@ -37,7 +38,7 @@ record Ideal {c : _} (pred : A → Set c) : Set (a ⊔ b ⊔ c) where
predicate = pred
generatedIdealPred : A A Set (a b)
generatedIdealPred a b = Sg A (λ c Setoid.__ S (a * c) b)
generatedIdealPred a b = a b
generatedIdeal : (a : A) Ideal (generatedIdealPred a)
Subgroup.isSubset (Ideal.isSubgroup (generatedIdeal a)) {x} {y} x=y (c , prC) = c , transitive prC x=y

View File

@@ -16,9 +16,13 @@ open import Rings.Lemmas
open import Sets.EquivalenceRelations
open import Rings.Ideals.Definition
open import Fields.Fields
open import Fields.Lemmas
open import Rings.Cosets
open import Rings.Ideals.Maximal.Definition
open import Rings.Ideals.Lemmas
open import Rings.Ideals.Prime.Definition
open import Rings.IntegralDomains.Definition
open import Rings.Ideals.Prime.Lemmas
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
@@ -77,3 +81,11 @@ MaximalIdeal.isMaximal (quotientFieldImpliesIdealMaximal f) {bigger} biggerIdeal
u = Ideal.closedUnderInverse biggerIdeal (Ideal.isSubset biggerIdeal *Commutative (Ideal.accumulatesTimes biggerIdeal biggerA))
v : bigger 1R
v = Ideal.isSubset biggerIdeal (invTwice additiveGroup 1R) (Ideal.closedUnderInverse biggerIdeal (Ideal.isSubset biggerIdeal (transitive (symmetric +Associative) (transitive (+WellDefined reflexive invRight) identRight)) (Ideal.closedUnderPlus biggerIdeal t u)))
idealMaximalImpliesIdealPrime : ({d : _} MaximalIdeal i {d}) PrimeIdeal i
idealMaximalImpliesIdealPrime max = quotientIntDomImpliesIdealPrime i f'
where
f : Field (cosetRing R i)
f = idealMaximalImpliesQuotientField max
f' : IntegralDomain (cosetRing R i)
f' = fieldIsIntDom f (λ p Field.nontrivial f (Equivalence.symmetric (Setoid.eq (cosetSetoid additiveGroup (Ideal.isSubgroup i))) p))

View File

@@ -9,6 +9,7 @@ open import Rings.Definition
open import Sets.EquivalenceRelations
open import Rings.Ideals.Definition
open import Rings.IntegralDomains.Definition
open import Rings.IntegralDomains.Lemmas
open import Rings.Ideals.Prime.Definition
open import Rings.Cosets
@@ -48,3 +49,18 @@ quotientIntDomImpliesIdealPrime intDom = record { isPrime = isPrime ; notContain
notCon 1=0 = IntegralDomain.nontrivial intDom (translate i 1=0)
isPrime : {a b : A} pred (a * b) (pred a False) pred b
isPrime {a} {b} predAB !predA = translate' i (IntegralDomain.intDom intDom (translate i predAB) λ t !predA (translate' i t))
private
dividesZero : {a : A} generatedIdealPred R 0R a a 0R
dividesZero (c , pr) = symmetric (transitive (symmetric (transitive *Commutative timesZero)) pr)
zeroIdealPrimeImpliesIntDom : PrimeIdeal (generatedIdeal R 0R) IntegralDomain R
IntegralDomain.intDom (zeroIdealPrimeImpliesIntDom record { isPrime = isPrime ; notContained = notContained ; notContainedIsNotContained = notContainedIsNotContained }) {a} {b} ab=0 a!=0 with isPrime {a} {b} (1R , transitive (transitive *Commutative timesZero) (symmetric ab=0)) λ 0|a a!=0 (dividesZero 0|a)
... | c , 0c=b = transitive (symmetric 0c=b) (transitive *Commutative timesZero)
IntegralDomain.nontrivial (zeroIdealPrimeImpliesIntDom record { isPrime = isPrime ; notContained = notContained ; notContainedIsNotContained = notContainedIsNotContained }) 1=0 = notContainedIsNotContained (notContained , transitive (*WellDefined (symmetric 1=0) reflexive) identIsIdent)
intDomImpliesZeroIdealPrime : IntegralDomain R PrimeIdeal (generatedIdeal R 0R)
PrimeIdeal.isPrime (intDomImpliesZeroIdealPrime intDom) (c , 0=ab) 0not|a with IntegralDomain.intDom intDom (transitive (symmetric 0=ab) (transitive *Commutative timesZero)) λ a=0 0not|a (0R , transitive timesZero (symmetric a=0))
... | b=0 = 0R , transitive timesZero (symmetric b=0)
PrimeIdeal.notContained (intDomImpliesZeroIdealPrime intDom) = 1R
PrimeIdeal.notContainedIsNotContained (intDomImpliesZeroIdealPrime intDom) (c , 0c=1) = IntegralDomain.nontrivial intDom (symmetric (transitive (symmetric (transitive *Commutative timesZero)) 0c=1))

View File

@@ -29,13 +29,3 @@ decidedIntDom : ({a b : A} → (a * b) (Ring.0R R) → (a 0R) || (b
decidedIntDom f ab=0 a!=0 with f ab=0
decidedIntDom f ab=0 a!=0 | inl x = exFalso (a!=0 x)
decidedIntDom f ab=0 a!=0 | inr x = x
cancelIntDom : (I : IntegralDomain) {a b c : A} (a * b) (a * c) ((a (Ring.0R R)) False) (b c)
cancelIntDom I {a} {b} {c} ab=ac a!=0 = transferToRight (Ring.additiveGroup R) t3
where
t1 : (a * b) + Group.inverse (Ring.additiveGroup R) (a * c) Ring.0R R
t1 = transferToRight'' (Ring.additiveGroup R) ab=ac
t2 : a * (b + Group.inverse (Ring.additiveGroup R) c) Ring.0R R
t2 = transitive (transitive (Ring.*DistributesOver+ R) (Group.+WellDefined (Ring.additiveGroup R) reflexive (transferToRight' (Ring.additiveGroup R) (transitive (symmetric (Ring.*DistributesOver+ R)) (transitive (Ring.*WellDefined R reflexive (Group.invLeft (Ring.additiveGroup R))) (Ring.timesZero R)))))) t1
t3 : (b + Group.inverse (Ring.additiveGroup R) c) Ring.0R R
t3 = IntegralDomain.intDom I t2 a!=0

View File

@@ -17,12 +17,3 @@ open import Fields.Fields
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
module Rings.IntegralDomains.Examples where
fieldIsIntDom : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ _*_ : A A A} {R : Ring S _+_ _*_} (F : Field R) (Setoid.__ S (Ring.1R R) (Ring.0R R) False) IntegralDomain R
IntegralDomain.intDom (fieldIsIntDom F 1!=0) {a} {b} ab=0 a!=0 with Field.allInvertible F a a!=0
IntegralDomain.intDom (fieldIsIntDom {S = S} {R = R} F _) {a} {b} ab=0 a!=0 | 1/a , prA = transitive (symmetric identIsIdent) (transitive (*WellDefined (symmetric prA) reflexive) (transitive (symmetric *Associative) (transitive (*WellDefined reflexive ab=0) timesZero)))
where
open Setoid S
open Equivalence eq
open Ring R
IntegralDomain.nontrivial (fieldIsIntDom F 1!=0) = 1!=0

View File

@@ -0,0 +1,32 @@
{-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae
open import Groups.Groups
open import Groups.Lemmas
open import Groups.Definition
open import Numbers.Naturals.Naturals
open import Orders
open import Setoids.Setoids
open import Functions
open import Rings.Definition
open import Rings.Lemmas
open import Sets.EquivalenceRelations
open import Rings.IntegralDomains.Definition
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
module Rings.IntegralDomains.Lemmas {m n : _} {A : Set n} {S : Setoid {n} {m} A} {_+_ _*_ : A A A} {R : Ring S _+_ _*_} (I : IntegralDomain R) where
open Setoid S
open Equivalence eq
open Ring R
cancelIntDom : {a b c : A} (a * b) (a * c) ((a (Ring.0R R)) False) (b c)
cancelIntDom {a} {b} {c} ab=ac a!=0 = transferToRight (Ring.additiveGroup R) t3
where
t1 : (a * b) + Group.inverse (Ring.additiveGroup R) (a * c) Ring.0R R
t1 = transferToRight'' (Ring.additiveGroup R) ab=ac
t2 : a * (b + Group.inverse (Ring.additiveGroup R) c) Ring.0R R
t2 = transitive (transitive (Ring.*DistributesOver+ R) (Group.+WellDefined (Ring.additiveGroup R) reflexive (transferToRight' (Ring.additiveGroup R) (transitive (symmetric (Ring.*DistributesOver+ R)) (transitive (Ring.*WellDefined R reflexive (Group.invLeft (Ring.additiveGroup R))) (Ring.timesZero R)))))) t1
t3 : (b + Group.inverse (Ring.additiveGroup R) c) Ring.0R R
t3 = IntegralDomain.intDom I t2 a!=0

View File

@@ -0,0 +1,29 @@
{-# OPTIONS --safe --warning=error --without-K #-}
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
open import Setoids.Setoids
open import Functions
open import Sets.EquivalenceRelations
open import Rings.Definition
open import Rings.Homomorphisms.Definition
open import Groups.Homomorphisms.Lemmas
open import Rings.IntegralDomains.Definition
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
module Rings.Irreducibles.Definition {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ _*_ : A A A} {R : Ring S _+_ _*_} (intDom : IntegralDomain R) where
open Setoid S
open Ring R
open import Rings.Units.Definition R
record Irreducible (r : A) : Set (a b) where
field
nonzero : (r 0R) False
nonunit : (Unit r) False
irreducible : (x y : A) (x * y) r (Unit x False) Unit y

View File

@@ -0,0 +1,30 @@
{-# OPTIONS --safe --warning=error --without-K #-}
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
open import Setoids.Setoids
open import Functions
open import Sets.EquivalenceRelations
open import Rings.Definition
open import Rings.Homomorphisms.Definition
open import Groups.Homomorphisms.Lemmas
open import Rings.IntegralDomains.Definition
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
module Rings.Primes.Definition {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ _*_ : A A A} {R : Ring S _+_ _*_} (intDom : IntegralDomain R) where
open import Rings.Divisible.Definition R
open Ring R
open Setoid S
open import Rings.Units.Definition R
record Prime (x : A) : Set (a b) where
field
isPrime : (r s : A) (x (r * s)) ((x r) False) (x s)
nonzero : (x 0R) False
nonunit : Unit x False

60
Rings/Primes/Lemmas.agda Normal file
View File

@@ -0,0 +1,60 @@
{-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae
open import Groups.Groups
open import Groups.Lemmas
open import Groups.Definition
open import Numbers.Naturals.Naturals
open import Orders
open import Setoids.Setoids
open import Functions
open import Rings.Definition
open import Rings.Lemmas
open import Sets.EquivalenceRelations
open import Fields.Fields
open import Rings.IntegralDomains.Definition
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
module Rings.Primes.Lemmas {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ _*_ : A A A} {R : Ring S _+_ _*_} (intDom : IntegralDomain R) where
open import Rings.Divisible.Definition R
open import Rings.IntegralDomains.Lemmas intDom
open import Rings.Ideals.Definition R
open import Rings.Primes.Definition intDom
open import Rings.Ideals.Prime.Definition {R = R}
open import Rings.Irreducibles.Definition intDom
open Ring R
open Setoid S
open Equivalence eq
primeImpliesPrimeIdeal : {a : A} Prime a PrimeIdeal (generatedIdeal a)
primeImpliesPrimeIdeal {a} record { isPrime = isPrime ; nonzero = nonzero ; nonunit = nonunit } = record { isPrime = λ {r} {s} isPrime r s ; notContained = 1R ; notContainedIsNotContained = bad }
where
bad : a 1R False
bad (c , ac=1) = nonunit (c , ac=1)
primeIdealImpliesPrime : {a : A} ((a 0R) False) PrimeIdeal (generatedIdeal a) Prime a
Prime.isPrime (primeIdealImpliesPrime {a} a!=0 record { isPrime = isPrime ; notContained = notContained ; notContainedIsNotContained = notContainedIsNotContained }) r s a|rs aNot|r = isPrime a|rs aNot|r
Prime.nonzero (primeIdealImpliesPrime {a} a!=0 record { isPrime = isPrime ; notContained = notContained ; notContainedIsNotContained = notContainedIsNotContained }) = a!=0
Prime.nonunit (primeIdealImpliesPrime {a} a!=0 record { isPrime = isPrime ; notContained = notContained ; notContainedIsNotContained = notContainedIsNotContained }) (c , ac=1) = notContainedIsNotContained ((c * notContained) , transitive *Associative (transitive (*WellDefined ac=1 reflexive) identIsIdent))
primeIsIrreducible : {a : A} Prime a Irreducible a
Irreducible.nonzero (primeIsIrreducible {a} prime) = Prime.nonzero prime
Irreducible.nonunit (primeIsIrreducible {a} prime) = Prime.nonunit prime
Irreducible.irreducible (primeIsIrreducible {a} prime) x y xy=a xNonunit = underlying pr , yUnit
where
a|xy : a (x * y)
a|xy = 1R , transitive *Commutative (transitive identIsIdent (symmetric xy=a))
a|yFalse : (a y) False
a|yFalse (c , ac=y) = xNonunit (c , transitive *Commutative t)
where
s : (a * (c * x)) a
s = transitive *Associative (transitive (*WellDefined ac=y reflexive) (transitive *Commutative xy=a))
t : (c * x) 1R
t = cancelIntDom {a} {c * x} {1R} (transitive s (symmetric (transitive *Commutative identIsIdent))) (Prime.nonzero prime)
pr : Sg A (λ c (a * c) x)
pr = Prime.isPrime prime y x (divisibleWellDefined reflexive *Commutative a|xy) a|yFalse
yUnit : (y * underlying pr) 1R
yUnit with pr
... | c , ac=x = transitive *Commutative (cancelIntDom {a} {c * y} {1R} (transitive *Associative (transitive (*WellDefined ac=x reflexive) (transitive xy=a (symmetric (transitive *Commutative identIsIdent))))) (Prime.nonzero prime))

View File

@@ -15,7 +15,7 @@ open import Groups.Homomorphisms.Lemmas
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
module Rings.PrincipalIdealDomain {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ _*_ : A A A} (R : Ring S _+_ _*_) where
module Rings.PrincipalIdealDomains.Definition {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ _*_ : A A A} (R : Ring S _+_ _*_) where
open import Rings.Ideals.Principal.Definition R
open import Rings.Ideals.Definition R

View File

@@ -15,6 +15,10 @@ open import Groups.Homomorphisms.Lemmas
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
module Rings.Ideals.Principal.Lemmas {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ _*_ : A A A} (R : Ring S _+_ _*_) where
module Rings.Units.Definition {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ _*_ : A A A} (R : Ring S _+_ _*_) where
fieldsArePid :
open Setoid S
open Ring R
Unit : A Set (a b)
Unit r = Sg A (λ s (r * s) 1R)

View File

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