{-# OPTIONS --safe --warning=error --without-K #-} open import LogicalFormulae open import Functions open import Groups.Groups open import Groups.Definition open import Groups.Lemmas open import Rings.Definition open import Setoids.Setoids open import Setoids.Orders open import Sets.EquivalenceRelations module Rings.Lemmas where ringTimesZero : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A → A → A} {_*_ : A → A → A} (R : Ring S _+_ _*_) → {x : A} → Setoid._∼_ S (x * (Ring.0R R)) (Ring.0R R) ringTimesZero {S = S} {_+_ = _+_} {_*_ = _*_} R {x} = symmetric (transitive blah'' (transitive (Group.multAssoc additiveGroup) (transitive (wellDefined invLeft reflexive) multIdentLeft))) where open Ring R open Group additiveGroup open Setoid S open Equivalence eq blah : (x * 0R) ∼ (x * 0R) + (x * 0R) blah = transitive (multWellDefined reflexive (symmetric multIdentRight)) multDistributes blah' : (inverse (x * 0R)) + (x * 0R) ∼ (inverse (x * 0R)) + ((x * 0R) + (x * 0R)) blah' = wellDefined reflexive blah blah'' : 0R ∼ (inverse (x * 0R)) + ((x * 0R) + (x * 0R)) blah'' = transitive (symmetric invLeft) blah' ringMinusExtracts : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A → A → A} {_*_ : A → A → A} (R : Ring S _+_ _*_) → {x y : A} → Setoid._∼_ S (x * Group.inverse (Ring.additiveGroup R) y) (Group.inverse (Ring.additiveGroup R) (x * y)) ringMinusExtracts {S = S} {_+_ = _+_} {_*_ = _*_} R {x = x} {y} = transferToRight' additiveGroup (transitive (symmetric multDistributes) (transitive (multWellDefined reflexive invLeft) (ringTimesZero R))) where open Setoid S open Equivalence eq open Ring R open Group additiveGroup ringAddInequalities : {a b c : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A → A → A} {_*_ : A → A → A} {R : Ring S _+_ _*_} {_<_ : Rel {_} {c} A} {pOrder : SetoidPartialOrder S _<_} {tOrder : SetoidTotalOrder pOrder} (oRing : OrderedRing R tOrder) → {w x y z : A} → w < x → y < z → (w + y) < (x + z) ringAddInequalities {S = S} {_+_ = _+_} {R = R} {_<_ = _<_} {pOrder = pOrder} oRing {w = w} {x} {y} {z} w