Rename some confusing fields (#51)

This commit is contained in:
Patrick Stevens
2019-10-13 09:31:54 +01:00
committed by GitHub
parent 96d15c6017
commit 959071214e
21 changed files with 441 additions and 441 deletions

View File

@@ -20,8 +20,8 @@ module Fields.FieldOfFractions where
fieldOfFractionsSetoid : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A A A} {_*_ : A A A} {R : Ring S _+_ _*_} (I : IntegralDomain R) Setoid (fieldOfFractionsSet I)
Setoid.__ (fieldOfFractionsSetoid {S = S} {_*_ = _*_} I) (a ,, (b , b!=0)) (c ,, (d , d!=0)) = Setoid.__ S (a * d) (b * c)
Equivalence.reflexive (Setoid.eq (fieldOfFractionsSetoid {R = R} I)) {a ,, (b , b!=0)} = Ring.multCommutative R
Equivalence.symmetric (Setoid.eq (fieldOfFractionsSetoid {S = S} {R = R} I)) {a ,, (b , b!=0)} {c ,, (d , d!=0)} ad=bc = transitive (Ring.multCommutative R) (transitive (symmetric ad=bc) (Ring.multCommutative R))
Equivalence.reflexive (Setoid.eq (fieldOfFractionsSetoid {R = R} I)) {a ,, (b , b!=0)} = Ring.*Commutative R
Equivalence.symmetric (Setoid.eq (fieldOfFractionsSetoid {S = S} {R = R} I)) {a ,, (b , b!=0)} {c ,, (d , d!=0)} ad=bc = transitive (Ring.*Commutative R) (transitive (symmetric ad=bc) (Ring.*Commutative R))
where
open Equivalence (Setoid.eq S)
Equivalence.transitive (Setoid.eq (fieldOfFractionsSetoid {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I)) {a ,, (b , b!=0)} {c ,, (d , d!=0)} {e ,, (f , f!=0)} ad=bc cf=de = p5
@@ -30,13 +30,13 @@ module Fields.FieldOfFractions where
open Ring R
open Equivalence eq
p : (a * d) * f (b * c) * f
p = Ring.multWellDefined R ad=bc reflexive
p = Ring.*WellDefined R ad=bc reflexive
p2 : (a * f) * d b * (d * e)
p2 = transitive (transitive (symmetric multAssoc) (transitive (multWellDefined reflexive multCommutative) multAssoc)) (transitive p (transitive (symmetric multAssoc) (multWellDefined reflexive cf=de)))
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 (multWellDefined reflexive multCommutative) multAssoc)
p3 = transitive p2 (transitive (*WellDefined reflexive *Commutative) *Associative)
p4 : (d 0R) || ((a * f) (b * e))
p4 = cancelIntDom I (transitive multCommutative (transitive p3 multCommutative))
p4 = cancelIntDom I (transitive *Commutative (transitive p3 *Commutative))
p5 : (a * f) (b * e)
p5 with p4
p5 | inl d=0 = exFalso (d!=0 d=0)
@@ -63,7 +63,7 @@ module Fields.FieldOfFractions where
ans pr | inr x = d!=0 x
fieldOfFractionsRing : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A A A} {_*_ : A A A} {R : Ring S _+_ _*_} (I : IntegralDomain R) Ring (fieldOfFractionsSetoid I) (fieldOfFractionsPlus I) (fieldOfFractionsTimes I)
Group.wellDefined (Ring.additiveGroup (fieldOfFractionsRing {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I)) {a ,, (b , b!=0)} {c ,, (d , d!=0)} {e ,, (f , f!=0)} {g ,, (h , h!=0)} af=be ch=dg = need
Group.+WellDefined (Ring.additiveGroup (fieldOfFractionsRing {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I)) {a ,, (b , b!=0)} {c ,, (d , d!=0)} {e ,, (f , f!=0)} {g ,, (h , h!=0)} af=be ch=dg = need
where
open Setoid S
open Ring R
@@ -73,79 +73,79 @@ module Fields.FieldOfFractions where
have2 : (a * f) (b * e)
have2 = af=be
need : (((a * d) + (b * c)) * (f * h)) ((b * d) * (((e * h) + (f * g))))
need = transitive (transitive (Ring.multCommutative R) (transitive (Ring.multDistributes R) (Group.wellDefined (Ring.additiveGroup R) (transitive multAssoc (transitive (multWellDefined (multCommutative) reflexive) (transitive (multWellDefined multAssoc reflexive) (transitive (multWellDefined (multWellDefined have2 reflexive) reflexive) (transitive (symmetric multAssoc) (transitive (multWellDefined reflexive multCommutative) (transitive multAssoc (transitive (multWellDefined (transitive (transitive (symmetric multAssoc) (multWellDefined reflexive multCommutative)) multAssoc) reflexive) (symmetric multAssoc))))))))) (transitive multCommutative (transitive (transitive (symmetric multAssoc) (multWellDefined reflexive (transitive (multWellDefined reflexive multCommutative) (transitive multAssoc (transitive (multWellDefined have1 reflexive) (transitive (symmetric multAssoc) (multWellDefined reflexive multCommutative))))))) multAssoc))))) (symmetric (Ring.multDistributes R))
Group.identity (Ring.additiveGroup (fieldOfFractionsRing {R = R} I)) = Ring.0R R ,, (Ring.1R R , IntegralDomain.nontrivial I)
need = transitive (transitive (Ring.*Commutative R) (transitive (Ring.*DistributesOver+ R) (Group.+WellDefined (Ring.additiveGroup R) (transitive *Associative (transitive (*WellDefined (*Commutative) reflexive) (transitive (*WellDefined *Associative reflexive) (transitive (*WellDefined (*WellDefined have2 reflexive) reflexive) (transitive (symmetric *Associative) (transitive (*WellDefined reflexive *Commutative) (transitive *Associative (transitive (*WellDefined (transitive (transitive (symmetric *Associative) (*WellDefined reflexive *Commutative)) *Associative) reflexive) (symmetric *Associative))))))))) (transitive *Commutative (transitive (transitive (symmetric *Associative) (*WellDefined reflexive (transitive (*WellDefined reflexive *Commutative) (transitive *Associative (transitive (*WellDefined have1 reflexive) (transitive (symmetric *Associative) (*WellDefined reflexive *Commutative))))))) *Associative))))) (symmetric (Ring.*DistributesOver+ R))
Group.0G (Ring.additiveGroup (fieldOfFractionsRing {R = R} I)) = Ring.0R R ,, (Ring.1R R , IntegralDomain.nontrivial I)
Group.inverse (Ring.additiveGroup (fieldOfFractionsRing {R = R} I)) (a ,, b) = Group.inverse (Ring.additiveGroup R) a ,, b
Group.multAssoc (Ring.additiveGroup (fieldOfFractionsRing {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I)) {a ,, (b , b!=0)} {c ,, (d , d!=0)} {e ,, (f , f!=0)} = need
Group.+Associative (Ring.additiveGroup (fieldOfFractionsRing {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I)) {a ,, (b , b!=0)} {c ,, (d , d!=0)} {e ,, (f , f!=0)} = need
where
open Setoid S
open Equivalence eq
need : (((a * (d * f)) + (b * ((c * f) + (d * e)))) * ((b * d) * f)) ((b * (d * f)) * ((((a * d) + (b * c)) * f) + ((b * d) * e)))
need = transitive (Ring.multCommutative R) (Ring.multWellDefined R (symmetric (Ring.multAssoc R)) (transitive (Group.wellDefined (Ring.additiveGroup R) reflexive (Ring.multDistributes R)) (transitive (Group.wellDefined (Ring.additiveGroup R) reflexive (Group.wellDefined (Ring.additiveGroup R) (Ring.multAssoc R) (Ring.multAssoc R))) (transitive (Group.multAssoc (Ring.additiveGroup R)) (Group.wellDefined (Ring.additiveGroup R) (transitive (transitive (Group.wellDefined (Ring.additiveGroup R) (transitive (Ring.multAssoc R) (Ring.multCommutative R)) (Ring.multCommutative R)) (symmetric (Ring.multDistributes R))) (Ring.multCommutative R)) reflexive)))))
Group.multIdentRight (Ring.additiveGroup (fieldOfFractionsRing {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I)) {a ,, (b , b!=0)} = need
need = transitive (Ring.*Commutative R) (Ring.*WellDefined R (symmetric (Ring.*Associative R)) (transitive (Group.+WellDefined (Ring.additiveGroup R) reflexive (Ring.*DistributesOver+ R)) (transitive (Group.+WellDefined (Ring.additiveGroup R) reflexive (Group.+WellDefined (Ring.additiveGroup R) (Ring.*Associative R) (Ring.*Associative R))) (transitive (Group.+Associative (Ring.additiveGroup R)) (Group.+WellDefined (Ring.additiveGroup R) (transitive (transitive (Group.+WellDefined (Ring.additiveGroup R) (transitive (Ring.*Associative R) (Ring.*Commutative R)) (Ring.*Commutative R)) (symmetric (Ring.*DistributesOver+ R))) (Ring.*Commutative R)) reflexive)))))
Group.identRight (Ring.additiveGroup (fieldOfFractionsRing {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I)) {a ,, (b , b!=0)} = need
where
open Setoid S
open Equivalence eq
need : (((a * Ring.1R R) + (b * Group.identity (Ring.additiveGroup R))) * b) ((b * Ring.1R R) * a)
need = transitive (transitive (Ring.multWellDefined R (transitive (Group.wellDefined (Ring.additiveGroup R) (transitive (Ring.multCommutative R) (Ring.multIdentIsIdent R)) reflexive) (transitive (Group.wellDefined (Ring.additiveGroup R) reflexive (ringTimesZero R)) (Group.multIdentRight (Ring.additiveGroup R)))) reflexive) (Ring.multCommutative R)) (symmetric (Ring.multWellDefined R (transitive (Ring.multCommutative R) (Ring.multIdentIsIdent R)) reflexive))
Group.multIdentLeft (Ring.additiveGroup (fieldOfFractionsRing {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I)) {a ,, (b , _)} = need
need : (((a * Ring.1R R) + (b * Group.0G (Ring.additiveGroup R))) * b) ((b * Ring.1R R) * a)
need = transitive (transitive (Ring.*WellDefined R (transitive (Group.+WellDefined (Ring.additiveGroup R) (transitive (Ring.*Commutative R) (Ring.identIsIdent R)) reflexive) (transitive (Group.+WellDefined (Ring.additiveGroup R) reflexive (Ring.timesZero R)) (Group.identRight (Ring.additiveGroup R)))) reflexive) (Ring.*Commutative R)) (symmetric (Ring.*WellDefined R (transitive (Ring.*Commutative R) (Ring.identIsIdent R)) reflexive))
Group.identLeft (Ring.additiveGroup (fieldOfFractionsRing {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I)) {a ,, (b , _)} = need
where
open Setoid S
open Equivalence eq
need : (((Group.identity (Ring.additiveGroup R) * b) + (Ring.1R R * a)) * b) ((Ring.1R R * b) * a)
need = transitive (transitive (Ring.multWellDefined R (transitive (Group.wellDefined (Ring.additiveGroup R) reflexive (Ring.multIdentIsIdent R)) (transitive (Group.wellDefined (Ring.additiveGroup R) (transitive (Ring.multCommutative R) (ringTimesZero R)) reflexive) (Group.multIdentLeft (Ring.additiveGroup R)))) reflexive) (Ring.multCommutative R)) (Ring.multWellDefined R (symmetric (Ring.multIdentIsIdent R)) reflexive)
need : (((Group.0G (Ring.additiveGroup R) * b) + (Ring.1R R * a)) * b) ((Ring.1R R * b) * a)
need = transitive (transitive (Ring.*WellDefined R (transitive (Group.+WellDefined (Ring.additiveGroup R) reflexive (Ring.identIsIdent R)) (transitive (Group.+WellDefined (Ring.additiveGroup R) (transitive (Ring.*Commutative R) (Ring.timesZero R)) reflexive) (Group.identLeft (Ring.additiveGroup R)))) reflexive) (Ring.*Commutative R)) (Ring.*WellDefined R (symmetric (Ring.identIsIdent R)) reflexive)
Group.invLeft (Ring.additiveGroup (fieldOfFractionsRing {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I)) {a ,, (b , _)} = need
where
open Setoid S
open Equivalence eq
need : (((Group.inverse (Ring.additiveGroup R) a * b) + (b * a)) * Ring.1R R) ((b * b) * Group.identity (Ring.additiveGroup R))
need = transitive (transitive (transitive (Ring.multCommutative R) (Ring.multIdentIsIdent R)) (transitive (Group.wellDefined (Ring.additiveGroup R) (Ring.multCommutative R) reflexive) (transitive (symmetric (Ring.multDistributes R)) (transitive (Ring.multWellDefined R reflexive (Group.invLeft (Ring.additiveGroup R))) (ringTimesZero R))))) (symmetric (ringTimesZero R))
need : (((Group.inverse (Ring.additiveGroup R) a * b) + (b * a)) * Ring.1R R) ((b * b) * Group.0G (Ring.additiveGroup R))
need = transitive (transitive (transitive (Ring.*Commutative R) (Ring.identIsIdent R)) (transitive (Group.+WellDefined (Ring.additiveGroup R) (Ring.*Commutative R) reflexive) (transitive (symmetric (Ring.*DistributesOver+ R)) (transitive (Ring.*WellDefined R reflexive (Group.invLeft (Ring.additiveGroup R))) (Ring.timesZero R))))) (symmetric (Ring.timesZero R))
Group.invRight (Ring.additiveGroup (fieldOfFractionsRing {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I)) {a ,, (b , _)} = need
where
open Setoid S
open Equivalence eq
need : (((a * b) + (b * Group.inverse (Ring.additiveGroup R) a)) * Ring.1R R) ((b * b) * Group.identity (Ring.additiveGroup R))
need = transitive (transitive (transitive (Ring.multCommutative R) (Ring.multIdentIsIdent R)) (transitive (Group.wellDefined (Ring.additiveGroup R) (Ring.multCommutative R) reflexive) (transitive (symmetric (Ring.multDistributes R)) (transitive (Ring.multWellDefined R reflexive (Group.invRight (Ring.additiveGroup R))) (ringTimesZero R))))) (symmetric (ringTimesZero R))
Ring.multWellDefined (fieldOfFractionsRing {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I) {a ,, (b , _)} {c ,, (d , _)} {e ,, (f , _)} {g ,, (h , _)} af=be ch=dg = need
need : (((a * b) + (b * Group.inverse (Ring.additiveGroup R) a)) * Ring.1R R) ((b * b) * Group.0G (Ring.additiveGroup R))
need = transitive (transitive (transitive (Ring.*Commutative R) (Ring.identIsIdent R)) (transitive (Group.+WellDefined (Ring.additiveGroup R) (Ring.*Commutative R) reflexive) (transitive (symmetric (Ring.*DistributesOver+ R)) (transitive (Ring.*WellDefined R reflexive (Group.invRight (Ring.additiveGroup R))) (Ring.timesZero R))))) (symmetric (Ring.timesZero R))
Ring.*WellDefined (fieldOfFractionsRing {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I) {a ,, (b , _)} {c ,, (d , _)} {e ,, (f , _)} {g ,, (h , _)} af=be ch=dg = need
where
open Setoid S
open Equivalence eq
need : ((a * c) * (f * h)) ((b * d) * (e * g))
need = transitive (Ring.multWellDefined R reflexive (Ring.multCommutative R)) (transitive (Ring.multAssoc R) (transitive (Ring.multWellDefined R (symmetric (Ring.multAssoc R)) reflexive) (transitive (Ring.multWellDefined R (Ring.multWellDefined R reflexive ch=dg) reflexive) (transitive (Ring.multCommutative R) (transitive (Ring.multAssoc R) (transitive (Ring.multWellDefined R (Ring.multCommutative R) reflexive) (transitive (Ring.multWellDefined R af=be reflexive) (transitive (Ring.multAssoc R) (transitive (Ring.multWellDefined R (transitive (symmetric (Ring.multAssoc R)) (transitive (Ring.multWellDefined R reflexive (Ring.multCommutative R)) (Ring.multAssoc R))) reflexive) (symmetric (Ring.multAssoc R)))))))))))
need = transitive (Ring.*WellDefined R reflexive (Ring.*Commutative R)) (transitive (Ring.*Associative R) (transitive (Ring.*WellDefined R (symmetric (Ring.*Associative R)) reflexive) (transitive (Ring.*WellDefined R (Ring.*WellDefined R reflexive ch=dg) reflexive) (transitive (Ring.*Commutative R) (transitive (Ring.*Associative R) (transitive (Ring.*WellDefined R (Ring.*Commutative R) reflexive) (transitive (Ring.*WellDefined R af=be reflexive) (transitive (Ring.*Associative R) (transitive (Ring.*WellDefined R (transitive (symmetric (Ring.*Associative R)) (transitive (Ring.*WellDefined R reflexive (Ring.*Commutative R)) (Ring.*Associative R))) reflexive) (symmetric (Ring.*Associative R)))))))))))
Ring.1R (fieldOfFractionsRing {R = R} I) = Ring.1R R ,, (Ring.1R R , IntegralDomain.nontrivial I)
Ring.groupIsAbelian (fieldOfFractionsRing {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I) {a ,, (b , _)} {c ,, (d , _)} = need
where
open Setoid S
open Equivalence eq
need : (((a * d) + (b * c)) * (d * b)) ((b * d) * ((c * b) + (d * a)))
need = transitive (Ring.multCommutative R) (Ring.multWellDefined R (Ring.multCommutative R) (transitive (Group.wellDefined (Ring.additiveGroup R) (Ring.multCommutative R) (Ring.multCommutative R)) (Ring.groupIsAbelian R)))
Ring.multAssoc (fieldOfFractionsRing {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I) {a ,, (b , _)} {c ,, (d , _)} {e ,, (f , _)} = need
need = transitive (Ring.*Commutative R) (Ring.*WellDefined R (Ring.*Commutative R) (transitive (Group.+WellDefined (Ring.additiveGroup R) (Ring.*Commutative R) (Ring.*Commutative R)) (Ring.groupIsAbelian R)))
Ring.*Associative (fieldOfFractionsRing {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I) {a ,, (b , _)} {c ,, (d , _)} {e ,, (f , _)} = need
where
open Setoid S
open Equivalence eq
need : ((a * (c * e)) * ((b * d) * f)) ((b * (d * f)) * ((a * c) * e))
need = transitive (Ring.multWellDefined R (Ring.multAssoc R) (symmetric (Ring.multAssoc R))) (Ring.multCommutative R)
Ring.multCommutative (fieldOfFractionsRing {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I) {a ,, (b , _)} {c ,, (d , _)} = need
need = transitive (Ring.*WellDefined R (Ring.*Associative R) (symmetric (Ring.*Associative R))) (Ring.*Commutative R)
Ring.*Commutative (fieldOfFractionsRing {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I) {a ,, (b , _)} {c ,, (d , _)} = need
where
open Setoid S
open Equivalence eq
need : ((a * c) * (d * b)) ((b * d) * (c * a))
need = transitive (Ring.multCommutative R) (Ring.multWellDefined R (Ring.multCommutative R) (Ring.multCommutative R))
Ring.multDistributes (fieldOfFractionsRing {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I) {a ,, (b , _)} {c ,, (d , _)} {e ,, (f , _)} = need
need = transitive (Ring.*Commutative R) (Ring.*WellDefined R (Ring.*Commutative R) (Ring.*Commutative R))
Ring.*DistributesOver+ (fieldOfFractionsRing {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I) {a ,, (b , _)} {c ,, (d , _)} {e ,, (f , _)} = need
where
open Setoid S
open Ring R
open Equivalence eq
inter : b * (a * ((c * f) + (d * e))) (((a * c) * (b * f)) + ((b * d) * (a * e)))
inter = transitive multAssoc (transitive multDistributes (Group.wellDefined additiveGroup (transitive multAssoc (transitive (multWellDefined (transitive (multWellDefined (multCommutative) reflexive) (transitive (symmetric multAssoc) (transitive (multWellDefined reflexive multCommutative) multAssoc))) reflexive) (symmetric multAssoc))) (transitive multAssoc (transitive (multWellDefined (transitive (symmetric multAssoc) (transitive (multWellDefined reflexive multCommutative) multAssoc)) reflexive) (symmetric multAssoc)))))
inter = transitive *Associative (transitive *DistributesOver+ (Group.+WellDefined additiveGroup (transitive *Associative (transitive (*WellDefined (transitive (*WellDefined (*Commutative) reflexive) (transitive (symmetric *Associative) (transitive (*WellDefined reflexive *Commutative) *Associative))) reflexive) (symmetric *Associative))) (transitive *Associative (transitive (*WellDefined (transitive (symmetric *Associative) (transitive (*WellDefined reflexive *Commutative) *Associative)) reflexive) (symmetric *Associative)))))
need : ((a * ((c * f) + (d * e))) * ((b * d) * (b * f))) ((b * (d * f)) * (((a * c) * (b * f)) + ((b * d) * (a * e))))
need = transitive (Ring.multWellDefined R reflexive (Ring.multWellDefined R reflexive (Ring.multCommutative R))) (transitive (Ring.multWellDefined R reflexive (Ring.multAssoc R)) (transitive (Ring.multCommutative R) (transitive (Ring.multWellDefined R (Ring.multWellDefined R (symmetric (Ring.multAssoc R)) reflexive) reflexive) (transitive (symmetric (Ring.multAssoc R)) (Ring.multWellDefined R reflexive inter)))))
Ring.multIdentIsIdent (fieldOfFractionsRing {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I) {a ,, (b , _)} = need
need = transitive (Ring.*WellDefined R reflexive (Ring.*WellDefined R reflexive (Ring.*Commutative R))) (transitive (Ring.*WellDefined R reflexive (Ring.*Associative R)) (transitive (Ring.*Commutative R) (transitive (Ring.*WellDefined R (Ring.*WellDefined R (symmetric (Ring.*Associative R)) reflexive) reflexive) (transitive (symmetric (Ring.*Associative R)) (Ring.*WellDefined R reflexive inter)))))
Ring.identIsIdent (fieldOfFractionsRing {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I) {a ,, (b , _)} = need
where
open Setoid S
open Equivalence eq
need : (((Ring.1R R) * a) * b) (((Ring.1R R * b)) * a)
need = transitive (Ring.multWellDefined R (Ring.multIdentIsIdent R) reflexive) (transitive (Ring.multCommutative R) (Ring.multWellDefined R (symmetric (Ring.multIdentIsIdent R)) reflexive))
need = transitive (Ring.*WellDefined R (Ring.identIsIdent R) reflexive) (transitive (Ring.*Commutative R) (Ring.*WellDefined R (symmetric (Ring.identIsIdent R)) reflexive))
fieldOfFractions : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A A A} {_*_ : A A A} {R : Ring S _+_ _*_} (I : IntegralDomain R) Field (fieldOfFractionsRing I)
Field.allInvertible (fieldOfFractions {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I) (fst ,, (b , _)) prA = (b ,, (fst , ans)) , need
@@ -153,13 +153,13 @@ module Fields.FieldOfFractions where
open Setoid S
open Equivalence eq
need : ((b * fst) * Ring.1R R) ((fst * b) * Ring.1R R)
need = Ring.multWellDefined R (Ring.multCommutative R) reflexive
need = Ring.*WellDefined R (Ring.*Commutative R) reflexive
ans : fst Ring.0R R False
ans pr = prA need'
where
need' : (fst * Ring.1R R) (b * Ring.0R R)
need' = transitive (Ring.multWellDefined R pr reflexive) (transitive (transitive (Ring.multCommutative R) (ringTimesZero R)) (symmetric (ringTimesZero R)))
Field.nontrivial (fieldOfFractions {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I) pr = IntegralDomain.nontrivial I (symmetric (transitive (symmetric (ringTimesZero R)) (transitive (Ring.multCommutative R) (transitive pr (Ring.multIdentIsIdent R)))))
need' = transitive (Ring.*WellDefined R pr reflexive) (transitive (transitive (Ring.*Commutative R) (Ring.timesZero R)) (symmetric (Ring.timesZero R)))
Field.nontrivial (fieldOfFractions {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I) pr = IntegralDomain.nontrivial I (symmetric (transitive (symmetric (Ring.timesZero R)) (transitive (Ring.*Commutative R) (transitive pr (Ring.identIsIdent R)))))
where
open Setoid S
open Equivalence eq
@@ -171,22 +171,22 @@ module Fields.FieldOfFractions where
homIntoFieldOfFractions : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A A A} {_*_ : A A A} {R : Ring S _+_ _*_} (I : IntegralDomain R) RingHom R (fieldOfFractionsRing I) (embedIntoFieldOfFractions I)
RingHom.preserves1 (homIntoFieldOfFractions {S = S} I) = Equivalence.reflexive (Setoid.eq S)
RingHom.ringHom (homIntoFieldOfFractions {S = S} {R = R} I) {a} {b} = Equivalence.transitive (Setoid.eq S) (Ring.multWellDefined R (Equivalence.reflexive (Setoid.eq S)) (Ring.multIdentIsIdent R)) (Ring.multCommutative R)
RingHom.ringHom (homIntoFieldOfFractions {S = S} {R = R} I) {a} {b} = Equivalence.transitive (Setoid.eq S) (Ring.*WellDefined R (Equivalence.reflexive (Setoid.eq S)) (Ring.identIsIdent R)) (Ring.*Commutative R)
GroupHom.groupHom (RingHom.groupHom (homIntoFieldOfFractions {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I)) {x} {y} = need
where
open Setoid S
open Equivalence eq
need : ((x + y) * (Ring.1R R * Ring.1R R)) (Ring.1R R * ((x * Ring.1R R) + (Ring.1R R * y)))
need = transitive (transitive (Ring.multWellDefined R reflexive (Ring.multIdentIsIdent R)) (transitive (Ring.multCommutative R) (transitive (Ring.multIdentIsIdent R) (Group.wellDefined (Ring.additiveGroup R) (symmetric (transitive (Ring.multCommutative R) (Ring.multIdentIsIdent R))) (symmetric (Ring.multIdentIsIdent R)))))) (symmetric (Ring.multIdentIsIdent R))
GroupHom.wellDefined (RingHom.groupHom (homIntoFieldOfFractions {S = S} {R = R} I)) x=y = transitive (Ring.multCommutative R) (Ring.multWellDefined R reflexive x=y)
need = transitive (transitive (Ring.*WellDefined R reflexive (Ring.identIsIdent R)) (transitive (Ring.*Commutative R) (transitive (Ring.identIsIdent R) (Group.+WellDefined (Ring.additiveGroup R) (symmetric (transitive (Ring.*Commutative R) (Ring.identIsIdent R))) (symmetric (Ring.identIsIdent R)))))) (symmetric (Ring.identIsIdent R))
GroupHom.wellDefined (RingHom.groupHom (homIntoFieldOfFractions {S = S} {R = R} I)) x=y = transitive (Ring.*Commutative R) (Ring.*WellDefined R reflexive x=y)
where
open Equivalence (Setoid.eq S)
homIntoFieldOfFractionsIsInj : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A A A} {_*_ : A A A} {R : Ring S _+_ _*_} (I : IntegralDomain R) SetoidInjection S (fieldOfFractionsSetoid I) (embedIntoFieldOfFractions I)
SetoidInjection.wellDefined (homIntoFieldOfFractionsIsInj {S = S} {R = R} I) x=y = transitive (Ring.multCommutative R) (Ring.multWellDefined R reflexive x=y)
SetoidInjection.wellDefined (homIntoFieldOfFractionsIsInj {S = S} {R = R} I) x=y = transitive (Ring.*Commutative R) (Ring.*WellDefined R reflexive x=y)
where
open Equivalence (Setoid.eq S)
SetoidInjection.injective (homIntoFieldOfFractionsIsInj {S = S} {R = R} I) x~y = transitive (symmetric multIdentIsIdent) (transitive multCommutative (transitive x~y multIdentIsIdent))
SetoidInjection.injective (homIntoFieldOfFractionsIsInj {S = S} {R = R} I) x~y = transitive (symmetric identIsIdent) (transitive *Commutative (transitive x~y identIsIdent))
where
open Ring R
open Setoid S