mirror of
https://github.com/Smaug123/agdaproofs
synced 2025-10-13 15:48:39 +00:00
Lots of rings (#82)
This commit is contained in:
@@ -6,7 +6,7 @@ open import Groups.Definition
|
||||
open import Groups.Lemmas
|
||||
open import Rings.Definition
|
||||
open import Rings.Lemmas
|
||||
open import Rings.IntegralDomains
|
||||
open import Rings.IntegralDomains.Definition
|
||||
open import Fields.Fields
|
||||
open import Functions
|
||||
open import Setoids.Setoids
|
||||
@@ -25,8 +25,7 @@ fieldOfFractionsPlus (a ,, (b , b!=0)) (c ,, (d , d!=0)) = (((a * d) + (b * c))
|
||||
open Ring R
|
||||
ans : ((b * d) ∼ Ring.0R R) → False
|
||||
ans pr with IntegralDomain.intDom I pr
|
||||
ans pr | inl x = b!=0 x
|
||||
ans pr | inr x = d!=0 x
|
||||
ans pr | f = exFalso (d!=0 (f b!=0))
|
||||
|
||||
plusWellDefined : {a b c d : fieldOfFractionsSet} → (Setoid._∼_ fieldOfFractionsSetoid a c) → (Setoid._∼_ fieldOfFractionsSetoid b d) → Setoid._∼_ fieldOfFractionsSetoid (fieldOfFractionsPlus a b) (fieldOfFractionsPlus c d)
|
||||
plusWellDefined {a ,, (b , b!=0)} {c ,, (d , d!=0)} {e ,, (f , f!=0)} {g ,, (h , h!=0)} af=be ch=dg = need
|
||||
|
@@ -6,7 +6,7 @@ open import Groups.Definition
|
||||
open import Groups.Lemmas
|
||||
open import Rings.Definition
|
||||
open import Rings.Lemmas
|
||||
open import Rings.IntegralDomains
|
||||
open import Rings.IntegralDomains.Definition
|
||||
open import Fields.Fields
|
||||
open import Functions
|
||||
open import Setoids.Setoids
|
||||
|
@@ -6,7 +6,7 @@ open import Groups.Definition
|
||||
open import Groups.Lemmas
|
||||
open import Rings.Definition
|
||||
open import Rings.Lemmas
|
||||
open import Rings.IntegralDomains
|
||||
open import Rings.IntegralDomains.Definition
|
||||
open import Fields.Fields
|
||||
open import Functions
|
||||
open import Setoids.Setoids
|
||||
|
@@ -8,7 +8,7 @@ open import Groups.Lemmas
|
||||
open import Rings.Definition
|
||||
open import Rings.Lemmas
|
||||
open import Rings.Homomorphisms.Definition
|
||||
open import Rings.IntegralDomains
|
||||
open import Rings.IntegralDomains.Definition
|
||||
open import Fields.Fields
|
||||
open import Functions
|
||||
open import Setoids.Setoids
|
||||
|
@@ -6,7 +6,7 @@ open import Groups.Definition
|
||||
open import Groups.Lemmas
|
||||
open import Rings.Definition
|
||||
open import Rings.Lemmas
|
||||
open import Rings.IntegralDomains
|
||||
open import Rings.IntegralDomains.Definition
|
||||
open import Fields.Fields
|
||||
open import Functions
|
||||
open import Setoids.Setoids
|
||||
@@ -25,8 +25,7 @@ fieldOfFractionsTimes (a ,, (b , b!=0)) (c ,, (d , d!=0)) = (a * c) ,, ((b * d)
|
||||
open Ring R
|
||||
ans : ((b * d) ∼ Ring.0R R) → False
|
||||
ans pr with IntegralDomain.intDom I pr
|
||||
ans pr | inl x = b!=0 x
|
||||
ans pr | inr x = d!=0 x
|
||||
ans pr | f = exFalso (d!=0 (f b!=0))
|
||||
|
||||
fieldOfFractionsTimesWellDefined : {a b c d : fieldOfFractionsSet} → (Setoid._∼_ fieldOfFractionsSetoid a c) → (Setoid._∼_ fieldOfFractionsSetoid b d) → (Setoid._∼_ fieldOfFractionsSetoid (fieldOfFractionsTimes a b) (fieldOfFractionsTimes c d))
|
||||
fieldOfFractionsTimesWellDefined {a ,, (b , _)} {c ,, (d , _)} {e ,, (f , _)} {g ,, (h , _)} af=be ch=dg = need
|
||||
|
File diff suppressed because it is too large
Load Diff
@@ -6,7 +6,7 @@ open import Groups.Definition
|
||||
open import Groups.Lemmas
|
||||
open import Rings.Definition
|
||||
open import Rings.Lemmas
|
||||
open import Rings.IntegralDomains
|
||||
open import Rings.IntegralDomains.Definition
|
||||
open import Fields.Fields
|
||||
open import Functions
|
||||
open import Setoids.Setoids
|
||||
|
@@ -6,16 +6,17 @@ open import Groups.Definition
|
||||
open import Groups.Lemmas
|
||||
open import Rings.Definition
|
||||
open import Rings.Lemmas
|
||||
open import Rings.IntegralDomains
|
||||
open import Fields.Fields
|
||||
open import Functions
|
||||
open import Setoids.Setoids
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Rings.IntegralDomains.Definition
|
||||
|
||||
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))))
|
||||
|
||||
@@ -36,9 +37,7 @@ Equivalence.transitive (Setoid.eq fieldOfFractionsSetoid) {a ,, (b , b!=0)} {c ,
|
||||
p2 = transitive (transitive (symmetric *Associative) (transitive (*WellDefined reflexive *Commutative) *Associative)) (transitive p (transitive (symmetric *Associative) (*WellDefined reflexive cf=de)))
|
||||
p3 : (a * f) * d ∼ (b * e) * d
|
||||
p3 = transitive p2 (transitive (*WellDefined reflexive *Commutative) *Associative)
|
||||
p4 : (d ∼ 0R) || ((a * f) ∼ (b * e))
|
||||
p4 = cancelIntDom I (transitive *Commutative (transitive p3 *Commutative))
|
||||
p4 : ((d ∼ 0R) → False) → ((a * f) ∼ (b * e))
|
||||
p4 = cancelIntDom R I (transitive *Commutative (transitive p3 *Commutative))
|
||||
p5 : (a * f) ∼ (b * e)
|
||||
p5 with p4
|
||||
p5 | inl d=0 = exFalso (d!=0 d=0)
|
||||
p5 | inr x = x
|
||||
p5 = p4 d!=0
|
||||
|
@@ -8,7 +8,7 @@ open import Rings.Lemmas
|
||||
open import Setoids.Setoids
|
||||
open import Setoids.Orders
|
||||
open import Orders
|
||||
open import Rings.IntegralDomains
|
||||
open import Rings.IntegralDomains.Definition
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
|
||||
@@ -23,6 +23,8 @@ record Field {m n : _} {A : Set m} {S : Setoid {m} {n} A} {_+_ : A → A → A}
|
||||
field
|
||||
allInvertible : (a : A) → ((a ∼ Group.0G (Ring.additiveGroup R)) → False) → Sg A (λ t → t * a ∼ 1R)
|
||||
nontrivial : (0R ∼ 1R) → False
|
||||
0F : A
|
||||
0F = Ring.0R R
|
||||
|
||||
record Field' {m n : _} : Set (lsuc m ⊔ lsuc n) where
|
||||
field
|
||||
|
@@ -11,7 +11,7 @@ open import Rings.Lemmas
|
||||
open import Setoids.Setoids
|
||||
open import Setoids.Orders
|
||||
open import Orders
|
||||
open import Rings.IntegralDomains
|
||||
open import Rings.IntegralDomains.Definition
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Fields.Fields
|
||||
@@ -84,35 +84,39 @@ abstract
|
||||
r : (1/2 * ((1R + 1R) * x)) ∼ (1/2' * ((1R + 1R) * x))
|
||||
r = Equivalence.transitive eq *Associative (Equivalence.transitive eq q (Equivalence.symmetric eq *Associative))
|
||||
|
||||
private
|
||||
orderedFieldIntDom : {a b : A} → (a * b ∼ 0R) → (a ∼ 0R) || (b ∼ 0R)
|
||||
orderedFieldIntDom {a} {b} ab=0 with totality 0R a
|
||||
... | inl (inl x) = inr (Equivalence.transitive eq (Equivalence.transitive eq (symmetric identIsIdent) (*WellDefined q reflexive)) p')
|
||||
where
|
||||
open Equivalence eq
|
||||
a!=0 : (a ∼ Group.0G additiveGroup) → False
|
||||
a!=0 pr = SetoidPartialOrder.irreflexive pOrder (SetoidPartialOrder.<WellDefined pOrder (symmetric pr) reflexive x)
|
||||
invA : A
|
||||
invA = underlying (Field.allInvertible F a a!=0)
|
||||
q : 1R ∼ (invA * a)
|
||||
q with Field.allInvertible F a a!=0
|
||||
... | invA , pr = symmetric pr
|
||||
p : invA * (a * b) ∼ invA * Group.0G additiveGroup
|
||||
p = *WellDefined reflexive ab=0
|
||||
p' : (invA * a) * b ∼ Group.0G additiveGroup
|
||||
p' = Equivalence.transitive eq (symmetric *Associative) (Equivalence.transitive eq p (Ring.timesZero R))
|
||||
orderedFieldIntDom {a} {b} ab=0 | inl (inr x) = inr (Equivalence.transitive eq (Equivalence.transitive eq (symmetric identIsIdent) (*WellDefined q reflexive)) p')
|
||||
where
|
||||
open Equivalence eq
|
||||
a!=0 : (a ∼ Group.0G additiveGroup) → False
|
||||
a!=0 pr = SetoidPartialOrder.irreflexive pOrder (SetoidPartialOrder.<WellDefined pOrder reflexive (symmetric pr) x)
|
||||
invA : A
|
||||
invA = underlying (Field.allInvertible F a a!=0)
|
||||
q : 1R ∼ (invA * a)
|
||||
q with Field.allInvertible F a a!=0
|
||||
... | invA , pr = symmetric pr
|
||||
p : invA * (a * b) ∼ invA * Group.0G additiveGroup
|
||||
p = *WellDefined reflexive ab=0
|
||||
p' : (invA * a) * b ∼ Group.0G additiveGroup
|
||||
p' = Equivalence.transitive eq (symmetric *Associative) (Equivalence.transitive eq p (Ring.timesZero R))
|
||||
orderedFieldIntDom {a} {b} ab=0 | inr x = inl (Equivalence.symmetric (Setoid.eq S) x)
|
||||
|
||||
orderedFieldIsIntDom : IntegralDomain R
|
||||
IntegralDomain.intDom orderedFieldIsIntDom {a} {b} ab=0 with totality (Ring.0R R) a
|
||||
IntegralDomain.intDom orderedFieldIsIntDom {a} {b} ab=0 | inl (inl x) = inr (Equivalence.transitive eq (Equivalence.transitive eq (symmetric identIsIdent) (*WellDefined q reflexive)) p')
|
||||
where
|
||||
open Equivalence eq
|
||||
a!=0 : (a ∼ Group.0G additiveGroup) → False
|
||||
a!=0 pr = SetoidPartialOrder.irreflexive pOrder (SetoidPartialOrder.<WellDefined pOrder (symmetric pr) reflexive x)
|
||||
invA : A
|
||||
invA = underlying (Field.allInvertible F a a!=0)
|
||||
q : 1R ∼ (invA * a)
|
||||
q with Field.allInvertible F a a!=0
|
||||
... | invA , pr = symmetric pr
|
||||
p : invA * (a * b) ∼ invA * Group.0G additiveGroup
|
||||
p = *WellDefined reflexive ab=0
|
||||
p' : (invA * a) * b ∼ Group.0G additiveGroup
|
||||
p' = Equivalence.transitive eq (symmetric *Associative) (Equivalence.transitive eq p (Ring.timesZero R))
|
||||
IntegralDomain.intDom orderedFieldIsIntDom {a} {b} ab=0 | inl (inr x) = inr (Equivalence.transitive eq (Equivalence.transitive eq (symmetric identIsIdent) (*WellDefined q reflexive)) p')
|
||||
where
|
||||
open Equivalence eq
|
||||
a!=0 : (a ∼ Group.0G additiveGroup) → False
|
||||
a!=0 pr = SetoidPartialOrder.irreflexive pOrder (SetoidPartialOrder.<WellDefined pOrder reflexive (symmetric pr) x)
|
||||
invA : A
|
||||
invA = underlying (Field.allInvertible F a a!=0)
|
||||
q : 1R ∼ (invA * a)
|
||||
q with Field.allInvertible F a a!=0
|
||||
... | invA , pr = symmetric pr
|
||||
p : invA * (a * b) ∼ invA * Group.0G additiveGroup
|
||||
p = *WellDefined reflexive ab=0
|
||||
p' : (invA * a) * b ∼ Group.0G additiveGroup
|
||||
p' = Equivalence.transitive eq (symmetric *Associative) (Equivalence.transitive eq p (Ring.timesZero R))
|
||||
IntegralDomain.intDom orderedFieldIsIntDom {a} {b} ab=0 | inr x = inl (Equivalence.symmetric (Setoid.eq S) x)
|
||||
IntegralDomain.intDom orderedFieldIsIntDom = decidedIntDom R orderedFieldIntDom
|
||||
IntegralDomain.nontrivial orderedFieldIsIntDom pr = Field.nontrivial F (Equivalence.symmetric (Setoid.eq S) pr)
|
||||
|
@@ -9,7 +9,6 @@ open import Rings.Lemmas
|
||||
open import Setoids.Setoids
|
||||
open import Setoids.Orders
|
||||
open import Orders
|
||||
open import Rings.IntegralDomains
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Fields.Fields
|
||||
|
@@ -10,7 +10,6 @@ open import Rings.Lemmas
|
||||
open import Setoids.Setoids
|
||||
open import Setoids.Orders
|
||||
open import Orders
|
||||
open import Rings.IntegralDomains
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Fields.Fields
|
||||
|
Reference in New Issue
Block a user