|
| 1 | +/- |
| 2 | +Copyright (c) 2018 Kenny Lau. All rights reserved. |
| 3 | +Released under Apache 2.0 license as described in the file LICENSE. |
| 4 | +Authors: Kenny Lau, Mario Carneiro, Johan Commelin, Amelia Livingston, Anne Baanen |
| 5 | +
|
| 6 | +! This file was ported from Lean 3 source module ring_theory.localization.inv_submonoid |
| 7 | +! leanprover-community/mathlib commit 6e7ca692c98bbf8a64868f61a67fb9c33b10770d |
| 8 | +! Please do not edit these lines, except to modify the commit id |
| 9 | +! if you have ported upstream changes. |
| 10 | +-/ |
| 11 | +import Mathlib.GroupTheory.Submonoid.Inverses |
| 12 | +import Mathlib.RingTheory.FiniteType |
| 13 | +import Mathlib.RingTheory.Localization.Basic |
| 14 | + |
| 15 | +/-! |
| 16 | +# Submonoid of inverses |
| 17 | +
|
| 18 | +## Main definitions |
| 19 | +
|
| 20 | + * `IsLocalization.invSubmonoid M S` is the submonoid of `S = M⁻¹R` consisting of inverses of |
| 21 | + each element `x ∈ M` |
| 22 | +
|
| 23 | +## Implementation notes |
| 24 | +
|
| 25 | +See `Mathlib/RingTheory/Localization/Basic.lean` for a design overview. |
| 26 | +
|
| 27 | +## Tags |
| 28 | +localization, ring localization, commutative ring localization, characteristic predicate, |
| 29 | +commutative ring, field of fractions |
| 30 | +-/ |
| 31 | + |
| 32 | + |
| 33 | +variable {R : Type _} [CommRing R] (M : Submonoid R) (S : Type _) [CommRing S] |
| 34 | + |
| 35 | +variable [Algebra R S] {P : Type _} [CommRing P] |
| 36 | + |
| 37 | +open Function |
| 38 | + |
| 39 | +open BigOperators |
| 40 | + |
| 41 | +namespace IsLocalization |
| 42 | + |
| 43 | +section InvSubmonoid |
| 44 | + |
| 45 | +/-- The submonoid of `S = M⁻¹R` consisting of `{ 1 / x | x ∈ M }`. -/ |
| 46 | +def invSubmonoid : Submonoid S := |
| 47 | + (M.map (algebraMap R S)).leftInv |
| 48 | +#align is_localization.inv_submonoid IsLocalization.invSubmonoid |
| 49 | + |
| 50 | +variable [IsLocalization M S] |
| 51 | + |
| 52 | +theorem submonoid_map_le_is_unit : M.map (algebraMap R S) ≤ IsUnit.submonoid S := by |
| 53 | + rintro _ ⟨a, ha, rfl⟩ |
| 54 | + exact IsLocalization.map_units S ⟨_, ha⟩ |
| 55 | +#align is_localization.submonoid_map_le_is_unit IsLocalization.submonoid_map_le_is_unit |
| 56 | + |
| 57 | +/-- There is an equivalence of monoids between the image of `M` and `invSubmonoid`. -/ |
| 58 | +noncomputable abbrev equivInvSubmonoid : M.map (algebraMap R S) ≃* invSubmonoid M S := |
| 59 | + ((M.map (algebraMap R S)).leftInvEquiv (submonoid_map_le_is_unit M S)).symm |
| 60 | +#align is_localization.equiv_inv_submonoid IsLocalization.equivInvSubmonoid |
| 61 | + |
| 62 | +/-- There is a canonical map from `M` to `invSubmonoid` sending `x` to `1 / x`. -/ |
| 63 | +noncomputable def toInvSubmonoid : M →* invSubmonoid M S := |
| 64 | + (equivInvSubmonoid M S).toMonoidHom.comp ((algebraMap R S : R →* S).submonoidMap M) |
| 65 | +#align is_localization.to_inv_submonoid IsLocalization.toInvSubmonoid |
| 66 | + |
| 67 | +theorem toInvSubmonoid_surjective : Function.Surjective (toInvSubmonoid M S) := |
| 68 | + Function.Surjective.comp (β := M.map (algebraMap R S)) |
| 69 | + (Equiv.surjective (equivInvSubmonoid _ _).toEquiv) (MonoidHom.submonoidMap_surjective _ _) |
| 70 | +#align is_localization.to_inv_submonoid_surjective IsLocalization.toInvSubmonoid_surjective |
| 71 | + |
| 72 | +@[simp] |
| 73 | +theorem toInvSubmonoid_mul (m : M) : (toInvSubmonoid M S m : S) * algebraMap R S m = 1 := |
| 74 | + Submonoid.leftInvEquiv_symm_mul _ (submonoid_map_le_is_unit _ _) _ |
| 75 | +#align is_localization.to_inv_submonoid_mul IsLocalization.toInvSubmonoid_mul |
| 76 | + |
| 77 | +@[simp] |
| 78 | +theorem mul_toInvSubmonoid (m : M) : algebraMap R S m * (toInvSubmonoid M S m : S) = 1 := |
| 79 | + Submonoid.mul_leftInvEquiv_symm _ (submonoid_map_le_is_unit _ _) ⟨_, _⟩ |
| 80 | +#align is_localization.mul_to_inv_submonoid IsLocalization.mul_toInvSubmonoid |
| 81 | + |
| 82 | +@[simp] |
| 83 | +theorem smul_toInvSubmonoid (m : M) : m • (toInvSubmonoid M S m : S) = 1 := by |
| 84 | + convert mul_toInvSubmonoid M S m |
| 85 | + ext |
| 86 | + rw [← Algebra.smul_def] |
| 87 | + rfl |
| 88 | +#align is_localization.smul_to_inv_submonoid IsLocalization.smul_toInvSubmonoid |
| 89 | + |
| 90 | +variable {S} |
| 91 | + |
| 92 | +-- Porting note: `surj'` was taken, so use `surj''` instead |
| 93 | +theorem surj'' (z : S) : ∃ (r : R) (m : M), z = r • (toInvSubmonoid M S m : S) := by |
| 94 | + rcases IsLocalization.surj M z with ⟨⟨r, m⟩, e : z * _ = algebraMap R S r⟩ |
| 95 | + refine' ⟨r, m, _⟩ |
| 96 | + rw [Algebra.smul_def, ← e, mul_assoc] |
| 97 | + simp |
| 98 | +#align is_localization.surj' IsLocalization.surj'' |
| 99 | + |
| 100 | +theorem toInvSubmonoid_eq_mk' (x : M) : (toInvSubmonoid M S x : S) = mk' S 1 x := by |
| 101 | + rw [← (IsLocalization.map_units S x).mul_left_inj] |
| 102 | + simp |
| 103 | +#align is_localization.to_inv_submonoid_eq_mk' IsLocalization.toInvSubmonoid_eq_mk' |
| 104 | + |
| 105 | +theorem mem_invSubmonoid_iff_exists_mk' (x : S) : |
| 106 | + x ∈ invSubmonoid M S ↔ ∃ m : M, mk' S 1 m = x := by |
| 107 | + simp_rw [← toInvSubmonoid_eq_mk'] |
| 108 | + exact ⟨fun h => ⟨_, congr_arg Subtype.val (toInvSubmonoid_surjective M S ⟨x, h⟩).choose_spec⟩, |
| 109 | + fun h => h.choose_spec ▸ (toInvSubmonoid M S h.choose).prop⟩ |
| 110 | +#align is_localization.mem_inv_submonoid_iff_exists_mk' IsLocalization.mem_invSubmonoid_iff_exists_mk' |
| 111 | + |
| 112 | +variable (S) |
| 113 | + |
| 114 | +set_option synthInstance.etaExperiment true in |
| 115 | +theorem span_invSubmonoid : Submodule.span R (invSubmonoid M S : Set S) = ⊤ := by |
| 116 | + rw [eq_top_iff] |
| 117 | + rintro x - |
| 118 | + rcases IsLocalization.surj'' M x with ⟨r, m, rfl⟩ |
| 119 | + exact Submodule.smul_mem _ _ (Submodule.subset_span (toInvSubmonoid M S m).prop) |
| 120 | +#align is_localization.span_inv_submonoid IsLocalization.span_invSubmonoid |
| 121 | + |
| 122 | +theorem finiteType_of_monoid_fg [Monoid.Fg M] : Algebra.FiniteType R S := by |
| 123 | + have := Monoid.fg_of_surjective _ (toInvSubmonoid_surjective M S) |
| 124 | + rw [Monoid.fg_iff_submonoid_fg] at this |
| 125 | + rcases this with ⟨s, hs⟩ |
| 126 | + refine' ⟨⟨s, _⟩⟩ |
| 127 | + rw [eq_top_iff] |
| 128 | + rintro x - |
| 129 | + change x ∈ (Subalgebra.toSubmodule (Algebra.adjoin R _ : Subalgebra R S) : Set S) |
| 130 | + rw [Algebra.adjoin_eq_span, hs, span_invSubmonoid] |
| 131 | + trivial |
| 132 | +#align is_localization.finite_type_of_monoid_fg IsLocalization.finiteType_of_monoid_fg |
| 133 | + |
| 134 | +end InvSubmonoid |
| 135 | + |
| 136 | +end IsLocalization |
0 commit comments