More lecture notes (#126)

This commit is contained in:
Patrick Stevens
2020-04-19 20:15:02 +01:00
committed by GitHub
parent 485b27e009
commit cab004f6d8
11 changed files with 163 additions and 36 deletions

View File

@@ -14,4 +14,7 @@ open import Sets.Cardinality.Infinite.Examples
open import ProjectEuler.Problem2
open import LectureNotes.NumbersAndSets.Examples1
open import LectureNotes.Groups.Lecture1
module Everything.Guardedness where

View File

@@ -150,7 +150,6 @@ open import Graphs.InducedSubgraph
open import LectureNotes.MetAndTop.Chapter1
open import LectureNotes.NumbersAndSets.Lecture1
open import LectureNotes.Groups.Lecture1
open import Modules.Examples
open import Modules.PolynomialModule

View File

@@ -8,6 +8,7 @@ open import Sets.EquivalenceRelations
open import LogicalFormulae
open import Rings.IntegralDomains.Definition
open import Rings.IntegralDomains.Lemmas
open import Setoids.Subset
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
@@ -30,13 +31,30 @@ abstract
halfHalves : {x : A} (1/2 : A) (pr : 1/2 + 1/2 1R) (x + x) * 1/2 x
halfHalves {x} 1/2 pr = transitive (transitive (transitive *Commutative (transitive (transitive *DistributesOver+ (transitive (+WellDefined *Commutative *Commutative) (symmetric *DistributesOver+))) *Commutative)) (*WellDefined pr (reflexive))) 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))))
IntegralDomain.nontrivial (fieldIsIntDom 1!=0) = 1!=0
fieldIsIntDom : IntegralDomain R
IntegralDomain.intDom fieldIsIntDom {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))))
IntegralDomain.nontrivial fieldIsIntDom 1=0 = Field.nontrivial F (symmetric 1=0)
allInvertibleWellDefined : {a b : A} {a!=0 : (a 0F) False} {b!=0 : (b 0F) False} (a b) underlying (allInvertible a a!=0) underlying (allInvertible b b!=0)
allInvertibleWellDefined {a} {b} {a!=0} {b!=0} a=b with allInvertible a a!=0
... | x , prX with allInvertible b b!=0
... | y , prY with transitive (transitive prX (symmetric prY)) (*WellDefined reflexive (symmetric a=b))
... | xa=ya = cancelIntDom (fieldIsIntDom λ p a!=0 (oneZeroImpliesAllZero (symmetric p))) (transitive *Commutative (transitive xa=ya *Commutative)) a!=0
... | xa=ya = cancelIntDom fieldIsIntDom (transitive *Commutative (transitive xa=ya *Commutative)) a!=0
private
mulNonzeros : Sg A (λ m (Setoid.__ S m (Ring.0R R)) False) Sg A (λ m (Setoid.__ S m (Ring.0R R)) False) Sg A (λ m (Setoid.__ S m (Ring.0R R)) False)
mulNonzeros (a , a!=0) (b , b!=0) = (a * b) , λ ab=0 b!=0 (IntegralDomain.intDom (fieldIsIntDom) ab=0 a!=0)
fieldMultiplicativeGroup : Group (subsetSetoid S {pred = λ m ((Setoid.__ S m (Ring.0R R)) False)}(λ {x} {y} x=y x!=0 λ y=0 x!=0 (Equivalence.transitive (Setoid.eq S) x=y y=0))) (mulNonzeros)
Group.+WellDefined (fieldMultiplicativeGroup) {x , prX} {y , prY} {z , prZ} {w , prW} = Ring.*WellDefined R
Group.0G (fieldMultiplicativeGroup) = Ring.1R R , λ 1=0 Field.nontrivial F (Equivalence.symmetric (Setoid.eq S) 1=0)
Group.inverse (fieldMultiplicativeGroup) (x , pr) with Field.allInvertible F x pr
... | 1/x , pr1/x = 1/x , λ 1/x=0 Field.nontrivial F (Equivalence.transitive (Setoid.eq S) (Equivalence.symmetric (Setoid.eq S) (Equivalence.transitive (Setoid.eq S) (Ring.*WellDefined R 1/x=0 (Equivalence.reflexive (Setoid.eq S))) (Ring.timesZero' R))) pr1/x)
Group.+Associative (fieldMultiplicativeGroup) {x , prX} {y , prY} {z , prZ} = Ring.*Associative R
Group.identRight (fieldMultiplicativeGroup) {x , prX} = Ring.identIsIdent' R
Group.identLeft (fieldMultiplicativeGroup) {x , prX} = Ring.identIsIdent R
Group.invLeft (fieldMultiplicativeGroup) {x , prX} with Field.allInvertible F x prX
... | 1/x , pr1/x = pr1/x
Group.invRight (fieldMultiplicativeGroup) {x , prX} with Field.allInvertible F x prX
... | 1/x , pr1/x = Equivalence.transitive (Setoid.eq S) (Ring.*Commutative R) pr1/x

View File

@@ -1,5 +1,6 @@
{-# OPTIONS --warning=error --safe --without-K #-}
{-# OPTIONS --warning=error --safe --without-K --guardedness #-}
open import Sets.EquivalenceRelations
open import Functions.Definition
open import Functions.Lemmas
open import Sets.FinSet.Definition
@@ -15,11 +16,17 @@ open import Numbers.Naturals.Semiring
open import Numbers.Naturals.Order
open import Numbers.Integers.Integers
open import Numbers.Rationals.Definition
open import Numbers.Reals.Definition
open import Rings.Definition
open import Fields.FieldOfFractions.Setoid
open import Semirings.Definition
open import Numbers.Modulo.Definition
open import Numbers.Modulo.Group
open import Fields.Fields
open import Setoids.Subset
open import Rings.IntegralDomains.Definition
open import Fields.Lemmas
open import Rings.Examples.Examples
module LectureNotes.Groups.Lecture1 where
@@ -28,9 +35,12 @@ module LectureNotes.Groups.Lecture1 where
groupExample1 : Group (reflSetoid ) (_+Z_)
groupExample1 = Group
groupExample2 : Group (fieldOfFractionsSetoid IntDom) (_+Q_)
groupExample2 : Group Setoid (_+Q_)
groupExample2 = Ring.additiveGroup Ring
groupExample2' : Group Setoid (_+R_)
groupExample2' = Ring.additiveGroup Ring
groupExample3 : Group (reflSetoid ) (_-Z_) False
groupExample3 record { +Associative = multAssoc } with multAssoc {nonneg 3} {nonneg 2} {nonneg 1}
groupExample3 record { +WellDefined = wellDefined } | ()
@@ -42,20 +52,14 @@ nonnegInjective : {a b : } → (nonneg a ≡ nonneg b) → a ≡ b
nonnegInjective {a} {.a} refl = refl
integersTimesNotGroup : Group (reflSetoid ) (_*Z_) False
integersTimesNotGroup record { +WellDefined = wellDefined ; 0G = (nonneg zero) ; inverse = inverse ; +Associative = multAssoc ; identRight = multIdentRight ; identLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } with multIdentLeft {negSucc 1}
... | ()
integersTimesNotGroup record { +WellDefined = wellDefined ; 0G = (nonneg (succ zero)) ; inverse = inverse ; +Associative = multAssoc ; identRight = multIdentRight ; identLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } with invLeft {nonneg zero}
... | bl with inverse (nonneg zero)
integersTimesNotGroup record { +WellDefined = wellDefined ; 0G = (nonneg (succ zero)) ; inverse = inverse ; +Associative = multAssoc ; identRight = multIdentRight ; identLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } | () | nonneg zero
integersTimesNotGroup record { +WellDefined = wellDefined ; 0G = (nonneg (succ zero)) ; inverse = inverse ; +Associative = multAssoc ; identRight = multIdentRight ; identLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } | p | nonneg (succ x) = naughtE (nonnegInjective (transitivity (applyEquality nonneg (equalityCommutative (Semiring.productZeroRight Semiring x))) p))
integersTimesNotGroup record { +WellDefined = wellDefined ; 0G = (nonneg (succ zero)) ; inverse = inverse ; +Associative = multAssoc ; identRight = multIdentRight ; identLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } | () | negSucc x
integersTimesNotGroup record { +WellDefined = wellDefined ; 0G = (nonneg (succ (succ x))) ; inverse = inverse ; +Associative = multAssoc ; identRight = multIdentRight ; identLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } with succInjective (negSuccInjective (multIdentLeft {negSucc 1}))
... | ()
integersTimesNotGroup record { +WellDefined = wellDefined ; 0G = (negSucc x) ; inverse = inverse ; +Associative = multAssoc ; identRight = multIdentRight ; identLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } with multIdentLeft {nonneg 2}
integersTimesNotGroup record { +WellDefined = wellDefined ; 0G = (negSucc x) ; inverse = inverse ; +Associative = multAssoc ; identRight = multIdentRight ; identLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } | ()
integersTimesNotGroup = multiplicationNotGroup Ring λ ()
rationalsTimesNotGroup : Group Setoid (_*Q_) False
rationalsTimesNotGroup = multiplicationNotGroup Ring λ ()
QNonzeroGroup : Group _ _
QNonzeroGroup = fieldMultiplicativeGroup Field
-- TODO: Q is not a group with *Q
-- TODO: Q without 0 is a group with *Q
-- TODO: {1, -1} is a group with *
nIsGroup : (n : ) (0<n : 0 <N n) _

View File

@@ -0,0 +1,97 @@
{-# OPTIONS --warning=error --safe --without-K --guardedness #-}
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import LogicalFormulae
open import Numbers.Integers.Integers
open import Numbers.Naturals.Semiring
open import Numbers.Naturals.Subtraction
open import Numbers.Naturals.Naturals
open import Numbers.Naturals.Order
open import Numbers.Naturals.Exponentiation
open import Numbers.Primes.PrimeNumbers
open import Maybe
open import Semirings.Definition
open import Numbers.Naturals.EuclideanAlgorithm
open import Sequences
open import Vectors
module LectureNotes.NumbersAndSets.Examples1 where
open Semiring Semiring
open import Semirings.Solver Semiring multiplicationNIsCommutative
private
abstract
q1Lemma1 : {q r : } 3 *N succ q +N r succ (succ (succ (3 *N q +N r)))
q1Lemma1 {q} {r} rewrite sumZeroRight q | commutative q (succ q) | commutative q (succ (succ (q +N q))) | +Associative q q q = refl
q1Lemma : {n : } (2 <N n) ((3 n) || (3 (n +N 2))) || (3 (n +N 4))
q1Lemma {zero} ()
q1Lemma {succ zero} (le zero ())
q1Lemma {succ zero} (le (succ x) ())
q1Lemma {succ (succ zero)} (le (succ x) proof) rewrite Semiring.commutative Semiring x 2 = exFalso (naughtE (equalityCommutative (succInjective (succInjective proof))))
q1Lemma {succ (succ (succ zero))} 2<n = inl (inl (aDivA 3))
q1Lemma {succ (succ (succ (succ zero)))} 2<n = inl (inr (divides (record { quot = 2 ; rem = zero ; pr = refl ; remIsSmall = inl (le 2 refl) ; quotSmall = inl (le 2 refl)}) refl))
q1Lemma {succ (succ (succ (succ (succ zero))))} 2<n = inr (divides (record { quot = 3 ; rem = zero ; pr = refl ; remIsSmall = inl (le 2 refl) ; quotSmall = inl (le 2 refl)}) refl)
q1Lemma {succ (succ (succ (succ (succ (succ n)))))} 2<n with q1Lemma {succ (succ (succ n))} (le n (applyEquality succ (Semiring.commutative Semiring n 2)))
q1Lemma {succ (succ (succ (succ (succ (succ n)))))} 2<n | inl (inl (divides record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = remIsSmall ; quotSmall = quotSmall } x)) = inl (inl (divides (record { quot = succ quot ; rem = rem ; pr = transitivity (q1Lemma1 {quot} {rem}) (applyEquality (λ x succ (succ (succ x))) pr) ; remIsSmall = remIsSmall ; quotSmall = inl (le 2 refl) }) x))
q1Lemma {succ (succ (succ (succ (succ (succ n)))))} 2<n | inl (inr (divides record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = remIsSmall ; quotSmall = quotSmall } x)) = inl (inr (divides record { quot = succ quot ; rem = rem ; pr = transitivity (q1Lemma1 {quot} {rem}) (applyEquality (λ x succ (succ (succ x))) pr) ; remIsSmall = remIsSmall ; quotSmall = inl (le 2 refl) } x))
q1Lemma {succ (succ (succ (succ (succ (succ n)))))} 2<n | inr (divides record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = remIsSmall ; quotSmall = quotSmall } x) = inr (divides record { quot = succ quot ; rem = rem ; pr = transitivity (q1Lemma1 {quot} {rem}) (applyEquality (λ x succ (succ (succ x))) pr) ; remIsSmall = remIsSmall ; quotSmall = inl (le 2 refl) } x)
q1 : {n : } Prime n Prime (n +N 2) Prime (n +N 4) n 3
q1 {succ zero} record { p>1 = (le zero ()) ; pr = pr } pN+2 pN+4
q1 {succ zero} record { p>1 = (le (succ x) ()) ; pr = pr } pN+2 pN+4
q1 {succ (succ zero)} pN record { p>1 = p>1 ; pr = pr } pN+4 with pr {2} (divides (record { quot = 2 ; rem = zero ; pr = refl ; remIsSmall = inl (le 1 refl) ; quotSmall = inl (le 1 refl)}) refl) (le 1 refl) (le 1 refl)
q1 {succ (succ zero)} pN record { p>1 = p>1 ; pr = pr } pN+4 | ()
q1 {succ (succ (succ n))} pN pN+2 pN+4 with q1Lemma {succ (succ (succ n))} (le n (applyEquality succ (commutative n 2)))
q1 {succ (succ (succ n))} pN pN+2 pN+4 | inl (inl 3|n) = equalityCommutative (primeDivPrimeImpliesEqual threeIsPrime pN 3|n)
q1 {succ (succ (succ n))} pN pN+2 pN+4 | inl (inr 3|n+2) with primeDivPrimeImpliesEqual threeIsPrime pN+2 3|n+2
... | bl rewrite commutative n 2 = exFalso (naughtE (succInjective (succInjective (succInjective bl))))
q1 {succ (succ (succ n))} pN pN+2 pN+4 | inr 3|n+4 with primeDivPrimeImpliesEqual threeIsPrime pN+4 3|n+4
... | bl rewrite commutative n 4 = exFalso (naughtE (succInjective (succInjective (succInjective bl))))
q3Differences :
q3Differences a = 2 *N a
q3Seq : (start : )
q3Seq start zero = start
q3Seq start (succ n) = q3Differences (succ n) +N q3Seq start n
q3SeqFirst : take 5 (funcToSequence (q3Seq 41)) 41 ,- 43 ,- 47 ,- 53 ,- 61 ,- []
q3SeqFirst = refl
q3N : (start : ) (a : ) q3Seq start a start +N (a +N (a *N a))
q3N start zero = equalityCommutative (sumZeroRight start)
q3N start (succ a) rewrite q3N start a | sumZeroRight a =
from (succ (plus (plus (const a) (succ (const a))) (plus (const start) (plus (const a) (times (const a) (const a))))))
to (plus (const start) (succ (plus (const a) (succ (plus (const a) (times (const a) (succ (const a))))))))
by applyEquality (λ i succ (succ i)) (transitivity (+Associative (a +N a) start _) (transitivity (transitivity (applyEquality (_+N (a +N a *N a)) (commutative (a +N a) start)) (equalityCommutative (+Associative start _ _))) (applyEquality (start +N_) (equalityCommutative (+Associative a a (a +N a *N a))))))
q3NDivision : (start : ) start (q3Seq start start)
q3NDivision 0 = aDivA 0
q3NDivision (succ start) rewrite q3N (succ start) (succ start) = dividesBothImpliesDividesSum {succ start} {succ start} (aDivA (succ start)) (dividesBothImpliesDividesSum {succ start} {succ start} (aDivA (succ start)) (divides (record { quot = succ start ; rem = 0 ; pr = sumZeroRight _ ; remIsSmall = inl (succIsPositive start) ; quotSmall = inl (succIsPositive start) }) refl))
q3 : (start : ) ((n : ) Prime (q3Seq start n)) False
q3 zero a with a 2
... | record { p>1 = p>1 ; pr = pr } with pr {2} (divides (record { quot = 3 ; rem = 0 ; pr = refl ; remIsSmall = inl (le 1 refl) ; quotSmall = inl (le 1 refl) }) refl) (le 3 refl) (succIsPositive 1)
... | ()
q3 (succ zero) a with a 0
... | record { p>1 = le zero () ; pr = pr }
... | record { p>1 = le (succ x) () ; pr = pr }
q3 (succ (succ start)) a with q3NDivision (succ (succ start))
... | r with a (succ (succ start))
... | s = compositeImpliesNotPrime (succ (succ start)) _ (succPreservesInequality (succIsPositive start)) (le (succ (succ start) +N ((start +N succ start) +N q3Seq (succ (succ start)) start)) (applyEquality (λ i succ (succ i)) (from (succ (plus (plus (const start) (plus (plus (const start) (succ (const start))) (const (q3Seq (succ (succ start)) start)))) (succ (succ (const start))))) to plus (plus (const start) (succ (succ (plus (const start) zero)))) (succ (plus (plus (const start) (succ (plus (const start) zero))) (const (q3Seq (succ (succ start)) start)))) by applyEquality (λ i succ (succ (succ (succ i)))) (transitivity (commutative _ start) (+Associative start start _))))) r s
q3' : ((n : ) Prime (q3Seq 41 n)) False
q3' = q3 41
q6 : {n : } 3 (n *N n) 3 n
q6 {n} 3|n^2 with primesArePrime {3} {n} {n} threeIsPrime 3|n^2
q6 {n} 3|n^2 | inl x = x
q6 {n} 3|n^2 | inr x = x
_*q10_ : {a : } {b : } (0 <N a) (0 <N b)
*q10Pos : {a : } {b : } (0<a : 0 <N a) (0<b : 0 <N b) 0 <N (_*q10_ 0<a 0<b)
_*q10_ {succ zero} {succ b} 0<a 0<b = succ (succ b)
_*q10_ {succ (succ a)} {succ zero} 0<a 0<b = _*q10_ {succ a} {2} (le a (applyEquality succ (sumZeroRight a))) (le 1 refl)
_*q10_ {succ (succ a)} {succ (succ b)} 0<a 0<b = _*q10_ {succ a} {_*q10_ {succ (succ a)} {succ b} 0<a (le b (applyEquality succ (sumZeroRight b)))} (le a (applyEquality succ (sumZeroRight a))) (*q10Pos 0<a (le b (applyEquality succ (sumZeroRight b))))
*q10Pos {succ zero} {succ b} 0<a 0<b = le (succ b) (applyEquality (λ i succ (succ i)) (sumZeroRight b))
*q10Pos {succ (succ a)} {succ zero} 0<a 0<b = *q10Pos (le a (applyEquality succ (sumZeroRight a))) (le 1 refl)
*q10Pos {succ (succ a)} {succ (succ b)} 0<a 0<b = *q10Pos (le a (applyEquality succ (sumZeroRight a))) (*q10Pos 0<a (le b (applyEquality succ (sumZeroRight b))))

View File

@@ -250,7 +250,6 @@ 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

View File

@@ -35,3 +35,5 @@ record Ring {n m} {A : Set n} (S : Setoid {n} {m} A) (_+_ : A → A → A) (_*_
*DistributesOver+' = transitive *Commutative (transitive *DistributesOver+ (+WellDefined *Commutative *Commutative))
*Associative' : {a b c : A} ((a * b) * c) (a * (b * c))
*Associative' = symmetric *Associative
identIsIdent' : {a : A} a * 1R a
identIsIdent' = transitive *Commutative identIsIdent

View File

@@ -14,7 +14,7 @@ open import Fields.Lemmas
module Rings.EuclideanDomains.Examples where
polynomialField : {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) EuclideanDomain R
EuclideanDomain.isIntegralDomain (polynomialField F 1!=0) = fieldIsIntDom F 1!=0
EuclideanDomain.isIntegralDomain (polynomialField F 1!=0) = fieldIsIntDom F
EuclideanDomain.norm (polynomialField F _) a!=0 = zero
EuclideanDomain.normSize (polynomialField F _) a!=0 b!=0 c b=ac = inr refl
DivisionAlgorithmResult.quotient (EuclideanDomain.divisionAlg (polynomialField {_*_ = _*_} F _) {a = a} {b} a!=0 b!=0) with Field.allInvertible F b b!=0

View File

@@ -10,17 +10,28 @@ open import Numbers.Modulo.Group
open import Numbers.Modulo.Definition
open import Rings.Examples.Proofs
open import Numbers.Primes.PrimeNumbers
open import Setoids.Setoids
open import Rings.Definition
open import Groups.Definition
open import Groups.Lemmas
open import Sets.EquivalenceRelations
module Rings.Examples.Examples where
nToZn : (n : ) (pr : 0 <N n) (x : ) n n pr
nToZn n pr x = nToZn' n pr x
multiplicationNotGroup : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ _*_ : A A A} (R : Ring S _+_ _*_) (nontrivial : Setoid.__ S (Ring.1R R) (Ring.0R R) False) Group S _*_ False
multiplicationNotGroup {S = S} R 1!=0 gr = exFalso (1!=0 (groupsHaveLeftCancellation gr (Ring.0R R) (Ring.1R R) (Ring.0R R) (transitive (Ring.timesZero' R) (symmetric (Ring.timesZero' R)))))
where
open Setoid S
open Equivalence eq
mod : (n : ) (pr : 0 <N n) n n pr
mod n pr a = mod' n pr a
nToZn : (n : ) (pr : 0 <N n) (x : ) n n pr
nToZn n pr x = nToZn' n pr x
modNExampleSurjective : (n : ) (pr : 0 <N n) Surjection (mod n pr)
modNExampleSurjective n pr = modNExampleSurjective' n pr
mod : (n : ) (pr : 0 <N n) n n pr
mod n pr a = mod' n pr a
modNExampleSurjective : (n : ) (pr : 0 <N n) Surjection (mod n pr)
modNExampleSurjective n pr = modNExampleSurjective' n pr
{-
modNExampleGroupHom : (n : ) → (pr : 0 <N n) → GroupHom Group (nGroup n pr) (mod n pr)

View File

@@ -80,12 +80,7 @@ MaximalIdeal.isMaximal (quotientFieldImpliesIdealMaximal f) {bigger} biggerIdeal
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))
idealMaximalImpliesIdealPrime max = quotientIntDomImpliesIdealPrime i (fieldIsIntDom (idealMaximalImpliesQuotientField max))
maximalIdealWellDefined : {d : _} {pred2 : A Set d} (i2 : Ideal R pred2) ({x : A} pred x pred2 x) ({x : A} pred2 x pred x) {e : _} MaximalIdeal i {e} MaximalIdeal i2 {e}
MaximalIdeal.notContained (maximalIdealWellDefined i2 pToP2 p2ToP record { notContained = notContained ; notContainedIsNotContained = notContainedIsNotContained ; isMaximal = isMaximal }) = notContained

View File

@@ -58,4 +58,3 @@ abstract
abelianUnderlyingGroup : AbelianGroup additiveGroup
abelianUnderlyingGroup = record { commutative = groupIsAbelian }