The Localization of a Commutative Ring

The Localization of a Commutative Ring

The Localization of a Commutative Ring Anthony Bordg February 23, 2021 Abstract We formalize the localization [1, II, §4] of a commutative ring R with respect to a multiplicative subset (i.e. a submonoid of R seen as a multiplicative monoid). This localization is itself a commutative ring and we build the natural homomorphism of rings from R to its localization. Contents 1 The Localization of a Commutative Ring 1 1.1 Localization ............................ 1 1.2 The Natural Homomorphism from a Ring to Its Localization 7 2 Acknowledgements 7 theory Localization imports Main HOL−Algebra:Group HOL−Algebra:Ring HOL−Algebra:AbelCoset begin Contents: • We define the localization of a commutative ring R with respect toa multiplicative subset, i.e. with respect to a submonoid of R (seen as a multiplicative monoid), cf. [rec-rng-of-frac]. • We prove that this localization is a commutative ring (cf. [crng-rng-of-frac]) equipped with a homomorphism of rings from R (cf. [rng-to-rng-of-frac-is-ring-hom]). 1 The Localization of a Commutative Ring 1.1 Localization locale submonoid = monoid M for M (structure) + fixes S assumes subset : S ⊆ carrier M 1 and m-closed [intro; simp] : [[x 2 S; y 2 S]] =) x ⊗ y 2 S and one-closed [simp]: 1 2 S lemma (in submonoid) is-submonoid: submonoid M S hproof i locale mult-submonoid-of-rng = ring R + submonoid R S for R and S locale mult-submonoid-of-crng = cring R + mult-submonoid-of-rng R S for R and S locale eq-obj-rng-of-frac = cring R + mult-submonoid-of-crng R S for R (structure) and S + fixes rel defines rel ≡ (jcarrier = carrier R × S; eq = λ(r;s)(r 0;s 0): 9 t2S: t ⊗ ((s 0⊗ r) (s ⊗ r 0)) = 0j) lemma (in abelian-group) minus-to-eq : assumes abelian-group G and x 2 carrier G and y 2 carrier G and x y = 0 shows x = y hproof i lemma (in eq-obj-rng-of-frac) equiv-obj-rng-of-frac: shows equivalence rel hproof i definition eq-class-of-rng-of-frac:: - ) 0a ) 0b ) -set (infix jı 10) 0 0 0 0 where r jrel s ≡ f(r ; s ) 2 carrier rel: (r; s) :=rel (r ; s )g lemma class-of-to-rel: shows class-of rel (r; s) = (r jrel s) hproof i lemma (in eq-obj-rng-of-frac) zero-in-mult-submonoid: assumes 0 2 S and (r; s) 2 carrier rel and (r 0; s 0) 2 carrier rel 0 0 shows (r jrel s) = (r jrel s ) hproof i definition set-eq-class-of-rng-of-frac:: - ) -set (set 0-class 0-of ı) where set-class-of rel ≡ f(r jrel s)j r s: (r; s) 2 carrier relg lemma elem-eq-class: assumes equivalence S and x 2 carrier S and y 2 carrier S and x :=S y shows class-of S x = class-of S y hproof i lemma (in abelian-group) four-elem-comm: assumes a 2 carrier G and b 2 carrier G and c 2 carrier G and d 2 carrier 2 G shows a c ⊕ b d = a ⊕ b c d hproof i lemma (in abelian-monoid) right-add-eq: assumes a = b shows c ⊕ a = c ⊕ b hproof i lemma (in abelian-monoid) right-minus-eq: assumes a = b shows c a = c b hproof i lemma (in abelian-group) inv-add: assumes a 2 carrier G and b 2 carrier G shows (a ⊕ b) = a b hproof i lemma (in abelian-group) right-inv-add: assumes a 2 carrier G and b 2 carrier G and c 2 carrier G shows c a b = c (a ⊕ b) hproof i context eq-obj-rng-of-frac begin definition carrier-rng-of-frac:: - partial-object where carrier-rng-of-frac ≡ (jcarrier = set-class-of relj) definition mult-rng-of-frac:: [-set; -set] ) -set where mult-rng-of-frac X Y ≡ let x 0 = (SOME x: x 2 X) in let y 0 = (SOME y: y 2 Y ) in 0 0 0 0 (fst x ⊗ fst y )jrel (snd x ⊗ snd y ) definition rec-monoid-rng-of-frac:: - monoid where rec-monoid-rng-of-frac ≡ (jcarrier = set-class-of rel; mult = mult-rng-of-frac; one = (1jrel 1)j) lemma member-class-to-carrier: 0 0 assumes x 2 (r jrel s) and y 2 (r jrel s ) shows (fst x ⊗ fst y; snd x ⊗ snd y) 2 carrier rel hproof i lemma member-class-to-member-class: 0 0 assumes x 2 (r jrel s) and y 2 (r jrel s ) shows (fst x ⊗ fst y jrel snd x ⊗ snd y) 2 set-class-of rel hproof i 3 lemma closed-mult-rng-of-frac : assumes (r; s) 2 carrier rel and (t; u) 2 carrier rel shows (r jrel s) ⊗rec-monoid-rng-of-frac (t jrel u) 2 set-class-of rel hproof i lemma non-empty-class: assumes (r; s) 2 carrier rel shows (r jrel s) 6= fg hproof i lemma mult-rng-of-frac-fundamental-lemma: assumes (r; s) 2 carrier rel and (r 0; s 0) 2 carrier rel 0 0 0 0 shows (r jrel s) ⊗rec-monoid-rng-of-frac (r jrel s ) = (r ⊗ r jrel s ⊗ s ) hproof i lemma member-class-to-assoc: assumes x 2 (r jrel s) and y 2 (t jrel u) and z 2 (v jrel w) shows ((fst x ⊗ fst y) ⊗ fst z jrel (snd x ⊗ snd y) ⊗ snd z) = (fst x ⊗ (fst y ⊗ fst z) jrel snd x ⊗ (snd y ⊗ snd z)) hproof i lemma assoc-mult-rng-of-frac: assumes (r; s) 2 carrier rel and (t; u) 2 carrier rel and (v; w) 2 carrier rel shows ((r jrel s) ⊗rec-monoid-rng-of-frac (t jrel u)) ⊗rec-monoid-rng-of-frac (v jrel w) = (r jrel s) ⊗rec-monoid-rng-of-frac ((t jrel u) ⊗rec-monoid-rng-of-frac (v jrel w)) hproof i lemma left-unit-mult-rng-of-frac: assumes (r; s) 2 carrier rel shows 1rec-monoid-rng-of-frac ⊗rec-monoid-rng-of-frac (r jrel s) = (r jrel s) hproof i lemma right-unit-mult-rng-of-frac: assumes (r; s) 2 carrier rel shows (r jrel s) ⊗rec-monoid-rng-of-frac 1rec-monoid-rng-of-frac = (r jrel s) hproof i lemma monoid-rng-of-frac: shows monoid (rec-monoid-rng-of-frac) hproof i lemma comm-mult-rng-of-frac: assumes (r; s) 2 carrier rel and (r 0; s 0) 2 carrier rel 0 0 0 0 shows (r jrel s) ⊗rec-monoid-rng-of-frac (r jrel s ) = (r jrel s ) ⊗rec-monoid-rng-of-frac (r jrel s) hproof i 4 lemma comm-monoid-rng-of-frac: shows comm-monoid (rec-monoid-rng-of-frac) hproof i definition add-rng-of-frac:: [-set; -set] ) -set where add-rng-of-frac X Y ≡ let x 0 = (SOME x: x 2 X) in let y 0 = (SOME y: y 2 Y ) in 0 0 0 0 0 0 (snd y ⊗ fst x ⊕ snd x ⊗ fst y ) jrel (snd x ⊗ snd y ) definition rec-rng-of-frac:: - ring where rec-rng-of-frac ≡ (jcarrier = set-class-of rel; mult = mult-rng-of-frac; one = (1jrel 1); zero = (0 jrel 1); add = add-rng-of-frac j) lemma add-rng-of-frac-fundamental-lemma: assumes (r; s) 2 carrier rel and (r 0; s 0) 2 carrier rel 0 0 0 0 0 shows (r jrel s) ⊕rec-rng-of-frac (r jrel s ) = (s ⊗ r ⊕ s ⊗ r jrel s ⊗ s ) hproof i lemma closed-add-rng-of-frac: assumes (r; s) 2 carrier rel and (r 0; s 0) 2 carrier rel 0 0 shows (r jrel s) ⊕rec-rng-of-frac (r jrel s ) 2 set-class-of rel hproof i lemma closed-rel-add: assumes (r; s) 2 carrier rel and (r 0; s 0) 2 carrier rel shows (s 0 ⊗ r ⊕ s ⊗ r 0; s ⊗ s 0) 2 carrier rel hproof i lemma assoc-add-rng-of-frac: assumes (r; s) 2 carrier rel and (r 0; s 0) 2 carrier rel and (r 00; s 00) 2 carrier rel 0 0 00 00 shows (r jrel s) ⊕rec-rng-of-frac (r jrel s ) ⊕rec-rng-of-frac (r jrel s ) = 0 0 00 00 (r jrel s) ⊕rec-rng-of-frac ((r jrel s ) ⊕rec-rng-of-frac (r jrel s )) hproof i lemma add-rng-of-frac-zero: shows (0 jrel 1) 2 set-class-of rel hproof i lemma l-unit-add-rng-of-frac: assumes (r; s) 2 carrier rel shows 0rec-rng-of-frac ⊕rec-rng-of-frac (r jrel s) = (r jrel s) hproof i lemma r-unit-add-rng-of-frac: assumes (r; s) 2 carrier rel shows (r jrel s) ⊕rec-rng-of-frac 0rec-rng-of-frac = (r jrel s) hproof i 5 lemma comm-add-rng-of-frac: assumes (r; s) 2 carrier rel and (r 0; s 0) 2 carrier rel 0 0 0 0 shows (r jrel s) ⊕rec-rng-of-frac (r jrel s ) = (r jrel s ) ⊕rec-rng-of-frac (r jrel s) hproof i lemma class-of-zero-rng-of-frac: assumes s 2 S shows (0 jrel s) = 0rec-rng-of-frac hproof i lemma r-inv-add-rng-of-frac: assumes (r; s) 2 carrier rel shows (r jrel s) ⊕rec-rng-of-frac ( r jrel s) = 0rec-rng-of-frac hproof i lemma l-inv-add-rng-of-frac: assumes (r; s) 2 carrier rel shows ( r jrel s) ⊕rec-rng-of-frac (r jrel s) = 0rec-rng-of-frac hproof i lemma abelian-group-rng-of-frac: shows abelian-group (rec-rng-of-frac) hproof i lemma r-distr-rng-of-frac: assumes (r; s) 2 carrier rel and (r 0; s 0) 2 carrier rel and (r 00; s 00) 2 carrier rel 0 0 00 00 shows ((r jrel s) ⊕rec-rng-of-frac (r jrel s )) ⊗rec-rng-of-frac (r jrel s ) = 00 00 0 0 (r jrel s) ⊗rec-rng-of-frac (r jrel s ) ⊕rec-rng-of-frac (r jrel s ) ⊗rec-rng-of-frac 00 00 (r jrel s ) hproof i lemma l-distr-rng-of-frac: assumes (r; s) 2 carrier rel and (r 0; s 0) 2 carrier rel and (r 00; s 00) 2 carrier rel 00 00 0 0 shows (r jrel s ) ⊗rec-rng-of-frac ((r jrel s) ⊕rec-rng-of-frac (r jrel s )) = 00 00 00 00 (r jrel s ) ⊗rec-rng-of-frac (r jrel s) ⊕rec-rng-of-frac (r jrel s ) ⊗rec-rng-of-frac 0 0 (r jrel s ) hproof i lemma rng-rng-of-frac: shows ring (rec-rng-of-frac) hproof i lemma crng-rng-of-frac: shows cring (rec-rng-of-frac) hproof i lemma simp-in-frac: 6 assumes (r; s) 2 carrier rel and s 0 2 S 0 0 shows (r jrel s) = (s ⊗ r jrel s ⊗ s) hproof i 1.2 The Natural Homomorphism from a Ring to Its Local- ization definition rng-to-rng-of-frac :: 0a ) ( 0a × 0a) set where rng-to-rng-of-frac r ≡ (r jrel 1) lemma rng-to-rng-of-frac-is-ring-hom : shows rng-to-rng-of-frac 2 ring-hom R rec-rng-of-frac hproof i lemma Im-rng-to-rng-of-frac-unit: assumes x 2 rng-to-rng-of-frac ‘ S shows x 2 Units rec-rng-of-frac hproof i lemma eq-class-to-rel: 0 0 assumes (r; s) 2 carrier R × S and (r ; s ) 2 carrier R × S and (r jrel s) = 0 0 (r jrel s ) 0 0 shows (r; s) :=rel (r ; s ) hproof i lemma rng-to-rng-of-frac-without-zero-div-is-inj: assumes 0 2= S and 8 a 2 carrier R:8 b 2 carrier R: a ⊗ b = 0 −! a = 0 _ b = 0 shows a-kernel R rec-rng-of-frac rng-to-rng-of-frac = f0g hproof i end end 2 Acknowledgements The author was supported by the ERC Advanced Grant ALEXANDRIA (Project 742178) funded by the European Research Council and led by Pro- fessor Lawrence Paulson at the University of Cambridge, UK.

View Full Text

Details

  • File Type
    pdf
  • Upload Time
    -
  • Content Languages
    English
  • Upload User
    Anonymous/Not logged-in
  • File Pages
    7 Page
  • File Size
    -

Download

Channel Download Status
Express Download Enable

Copyright

We respect the copyrights and intellectual property rights of all users. All uploaded documents are either original works of the uploader or authorized works of the rightful owners.

  • Not to be reproduced or distributed without explicit permission.
  • Not used for commercial purposes outside of approved use cases.
  • Not used to infringe on the rights of the original creators.
  • If you believe any content infringes your copyright, please contact us immediately.

Support

For help with questions, suggestions, or problems, please contact us