Split out parts of Field of Fractions (#63)

This commit is contained in:
Patrick Stevens
2019-11-02 21:31:46 +00:00
committed by GitHub
parent 1325236359
commit e4daab7153
12 changed files with 374 additions and 224 deletions

View File

@@ -0,0 +1,44 @@
{-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae
open import Groups.Groups
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 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))))
fieldOfFractionsSetoid : Setoid fieldOfFractionsSet
Setoid.__ fieldOfFractionsSetoid (a ,, (b , b!=0)) (c ,, (d , d!=0)) = Setoid.__ S (a * d) (b * c)
Equivalence.reflexive (Setoid.eq fieldOfFractionsSetoid) {a ,, (b , b!=0)} = Ring.*Commutative R
Equivalence.symmetric (Setoid.eq fieldOfFractionsSetoid) {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) {a ,, (b , b!=0)} {c ,, (d , d!=0)} {e ,, (f , f!=0)} ad=bc cf=de = p5
where
open Setoid S
open Ring R
open Equivalence eq
p : (a * d) * f (b * c) * f
p = Ring.*WellDefined R ad=bc reflexive
p2 : (a * f) * d b * (d * e)
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))
p5 : (a * f) (b * e)
p5 with p4
p5 | inl d=0 = exFalso (d!=0 d=0)
p5 | inr x = x