Files
agdaproofs/Groups/Examples/Examples.agda
2019-12-07 13:00:18 +00:00

140 lines
9.4 KiB
Agda
Raw Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

{-# OPTIONS --safe --warning=error --with-K #-}
open import Sets.EquivalenceRelations
open import Groups.Definition
open import Orders
open import Rings.Definition
open import Numbers.Integers.Integers
open import Numbers.Integers.Multiplication
open import Setoids.Setoids
open import LogicalFormulae
open import Sets.FinSet
open import Functions
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
open import Numbers.Modulo.Group
open import Rings.Examples.Examples
open import Numbers.Primes.PrimeNumbers
open import Groups.Lemmas
open import Groups.Homomorphisms.Definition
open import Groups.Homomorphisms.Lemmas
open import Groups.Isomorphisms.Definition
open import Groups.Cyclic.Definition
open import Groups.QuotientGroup.Definition
open import Groups.Subgroups.Definition
open import Groups.Subgroups.Normal.Definition
module Groups.Examples.Examples where
trivialGroup : Group (reflSetoid (FinSet 1)) λ _ _ fzero
Group.+WellDefined trivialGroup _ _ = refl
Group.0G trivialGroup = fzero
Group.inverse trivialGroup _ = fzero
Group.+Associative trivialGroup = refl
Group.identRight trivialGroup {fzero} = refl
Group.identRight trivialGroup {fsucc ()}
Group.identLeft trivialGroup {fzero} = refl
Group.identLeft trivialGroup {fsucc ()}
Group.invLeft trivialGroup = refl
Group.invRight trivialGroup = refl
elementPowZ : (n : ) (elementPower Group (nonneg 1) (nonneg n)) nonneg n
elementPowZ zero = refl
elementPowZ (succ n) rewrite elementPowZ n = refl
Cyclic : CyclicGroup Group
CyclicGroup.generator Cyclic = nonneg 1
CyclicGroup.cyclic Cyclic {nonneg x} = (nonneg x , elementPowZ x)
CyclicGroup.cyclic Cyclic {negSucc x} = (negSucc x , ans)
where
ans : (Group.inverse Group ((nonneg 1) +Z elementPower Group (nonneg 1) (nonneg x))) negSucc x
ans rewrite elementPowZ x = refl
elementPowZn : (n : ) {pr : 0 <N succ (succ n)} (power : ) (powerLess : power <N succ (succ n)) {p : 1 <N succ (succ n)} elementPower (nGroup (succ (succ n)) pr) (record { x = 1 ; xLess = p }) (nonneg power) record { x = power ; xLess = powerLess }
elementPowZn n zero powerLess = equalityZn _ _ refl
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 }
t = elementPowZn n {pr} power (PartialOrder.<Transitive (TotalOrder.order TotalOrder) (a<SuccA power) powerLess) {succPreservesInequality (succIsPositive n)}
u : elementPower (nGroup (succ (succ n)) pr) (record { x = 1 ; xLess = p }) (nonneg power) record { x = power ; xLess = PartialOrder.<Transitive (TotalOrder.order TotalOrder) (a<SuccA power) powerLess }
u rewrite <NRefl p (succPreservesInequality (succIsPositive n)) = t
v : n.x (elementPower (nGroup (succ (succ n)) pr) (record { x = 1 ; xLess = p }) (nonneg power)) power
v rewrite u = refl
elementPowZn n {pr} (succ power) powerLess {p} | inl (inr x) with elementPower (nGroup (succ (succ n)) pr) (record { x = 1 ; xLess = p }) (nonneg power)
elementPowZn n {pr} (succ power) powerLess {p} | inl (inr x) | record { x = x' ; xLess = xLess } = exFalso (noIntegersBetweenXAndSuccX _ x xLess)
elementPowZn n {pr} (succ power) powerLess {p} | inr x with inspect (elementPower (nGroup (succ (succ n)) pr) (record { x = 1 ; xLess = p }) (nonneg power))
elementPowZn n {pr} (succ power) powerLess {p} | inr x | record { x = x' ; xLess = p' } with inspected rewrite elementPowZn n {pr} power (PartialOrder.<Transitive (TotalOrder.order TotalOrder) (a<SuccA power) powerLess) {p} | x = exFalso (PartialOrder.irreflexive (TotalOrder.order TotalOrder) powerLess)
nCyclic : (n : ) (pr : 0 <N n) CyclicGroup (nGroup n pr)
nCyclic zero ()
CyclicGroup.generator (nCyclic (succ zero) pr) = record { x = 0 ; xLess = pr }
CyclicGroup.cyclic (nCyclic (succ zero) pr) {record { x = zero ; xLess = xLess }} rewrite <NRefl xLess (succIsPositive 0) | <NRefl pr (succIsPositive 0) = (nonneg 0 , refl)
CyclicGroup.cyclic (nCyclic (succ zero) _) {record { x = succ x ; xLess = xLess }} = exFalso (zeroNeverGreater (canRemoveSuccFrom<N xLess))
nCyclic (succ (succ n)) pr = record { generator = record { x = 1 ; xLess = succPreservesInequality (succIsPositive n) } ; cyclic = λ {x} ans x }
where
ans : (z : n (succ (succ n)) pr) Sg (λ i (elementPower (nGroup (succ (succ n)) pr) (record { x = 1 ; xLess = succPreservesInequality (succIsPositive n) }) i) z)
ans record { x = x ; xLess = xLess } = nonneg x , elementPowZn n x xLess
multiplyByNHom : (n : ) GroupHom Group Group (λ i i *Z nonneg n)
GroupHom.groupHom (multiplyByNHom n) {x} {y} = lemma1
where
open Setoid (reflSetoid )
open Group Group
lemma : nonneg n *Z (x +Z y) (nonneg n) *Z x +Z nonneg n *Z y
lemma = Ring.*DistributesOver+ Ring {nonneg n} {x} {y}
lemma1 : (x +Z y) *Z nonneg n x *Z nonneg n +Z y *Z nonneg n
lemma1 rewrite Ring.*Commutative Ring {x +Z y} {nonneg n} | Ring.*Commutative Ring {x} {nonneg n} | Ring.*Commutative Ring {y} {nonneg n} = lemma
GroupHom.wellDefined (multiplyByNHom n) {x} {.x} refl = refl
modNExampleGroupHom : (n : ) (pr : 0 <N n) GroupHom Group (nGroup n pr) (mod n pr)
GroupHom.groupHom (modNExampleGroupHom n 0<n) {x} {y} = {!!}
GroupHom.wellDefined (modNExampleGroupHom n 0<n) {x} {.x} refl = refl
intoZn : {n : } {pr : 0 <N n} (x : ) (x<n : x <N n) mod n pr (nonneg x) record { x = x ; xLess = x<n }
intoZn {zero} {()}
intoZn {succ n} {0<n} x x<n with divisionAlg (succ n) x
intoZn {succ n} {0<n} x x<n | record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = inl y ; quotSmall = quotSmall } = equalityZn _ _ (modIsUnique (record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = inl y ; quotSmall = quotSmall }) (record { quot = 0 ; rem = x ; pr = ans ; remIsSmall = inl x<n ; quotSmall = inl (succIsPositive n)}))
where
ans : n *N 0 +N x x
ans rewrite multiplicationNIsCommutative n 0 = refl
intoZn {succ n} {0<n} x x<n | record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = inr () ; quotSmall = quotSmall }
quotientZn : (n : ) (pr : 0 <N n) GroupIso (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
given : mod n pr (x +Z (Group.inverse Group y)) record { x = 0 ; xLess = pr }
given = x~y
given' : (mod n pr x) +n (mod n pr (Group.inverse Group y)) record { x = 0 ; xLess = pr }
given' = transitivity (equalityCommutative (GroupHom.groupHom (modNExampleGroupHom n pr))) given
given'' : (mod n pr x) +n Group.inverse (nGroup n pr) (mod n pr y) record { x = 0 ; xLess = pr}
given'' = transitivity (applyEquality (λ i mod n pr x +n i) (equalityCommutative (homRespectsInverse (modNExampleGroupHom n pr)))) given'
need : mod n pr x mod n pr y
need = transferToRight (nGroup n pr) given''
SetoidInjection.wellDefined (SetoidBijection.inj (GroupIso.bij (quotientZn n pr))) {x} {y} x=y = {!!}
SetoidSurjection.wellDefined (SetoidBijection.surj (GroupIso.bij (quotientZn n pr))) {x} {y} = SetoidInjection.wellDefined (SetoidBijection.inj (GroupIso.bij (quotientZn n pr))) {x} {y}
SetoidInjection.injective (SetoidBijection.inj (GroupIso.bij (quotientZn n pr))) {x} {y} fx~fy = need
where
given : mod n pr x +n Group.inverse (nGroup n pr) (mod n pr y) record { x = 0 ; xLess = pr }
given = transferToRight'' (nGroup n pr) fx~fy
given' : mod n pr (x +Z (Group.inverse Group y)) record { x = 0 ; xLess = pr }
given' = identityOfIndiscernablesLeft _≡_ given (equalityCommutative (transitivity (GroupHom.groupHom (modNExampleGroupHom n pr) {x}) (applyEquality (λ i mod n pr x +n i) (homRespectsInverse (modNExampleGroupHom n pr)))))
need' : (mod n pr x) +n (mod n pr (Group.inverse Group y)) record { x = 0 ; xLess = pr }
need' rewrite equalityCommutative (GroupHom.groupHom (modNExampleGroupHom n pr) {x} {Group.inverse Group y}) = given'
need : mod n pr (x +Z (Group.inverse Group y)) record { x = 0 ; xLess = pr}
need = identityOfIndiscernablesLeft _≡_ need' (equalityCommutative (GroupHom.groupHom (modNExampleGroupHom n pr)))
SetoidSurjection.surjective (SetoidBijection.surj (GroupIso.bij (quotientZn n pr))) {record { x = x ; xLess = xLess }} = nonneg x , intoZn x xLess
trivialGroupHom : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A A A} (G : Group S _+_) GroupHom trivialGroup G (λ _ Group.0G G)
GroupHom.groupHom (trivialGroupHom {S = S} G) = symmetric identRight
where
open Setoid S
open Group G
open Equivalence eq
GroupHom.wellDefined (trivialGroupHom {S = S} G) _ = Equivalence.reflexive (Setoid.eq S)