|
| 1 | +/- |
| 2 | +Copyright (c) 2021 Andrew Yang. All rights reserved. |
| 3 | +Released under Apache 2.0 license as described in the file LICENSE. |
| 4 | +Authors: Andrew Yang |
| 5 | +
|
| 6 | +! This file was ported from Lean 3 source module ring_theory.ring_hom.finite_type |
| 7 | +! leanprover-community/mathlib commit 64fc7238fb41b1a4f12ff05e3d5edfa360dd768c |
| 8 | +! Please do not edit these lines, except to modify the commit id |
| 9 | +! if you have ported upstream changes. |
| 10 | +-/ |
| 11 | +import Mathlib.RingTheory.LocalProperties |
| 12 | +import Mathlib.RingTheory.Localization.InvSubmonoid |
| 13 | + |
| 14 | +/-! |
| 15 | +
|
| 16 | +# The meta properties of finite-type ring homomorphisms. |
| 17 | +
|
| 18 | +The main result is `RingHom.finiteType_is_local`. |
| 19 | +
|
| 20 | +-/ |
| 21 | + |
| 22 | + |
| 23 | +namespace RingHom |
| 24 | + |
| 25 | +open scoped Pointwise |
| 26 | + |
| 27 | +theorem finiteType_stableUnderComposition : StableUnderComposition @FiniteType := by |
| 28 | + introv R hf hg |
| 29 | + exact hg.comp hf |
| 30 | +#align ring_hom.finite_type_stable_under_composition RingHom.finiteType_stableUnderComposition |
| 31 | + |
| 32 | +theorem finiteType_holdsForLocalizationAway : HoldsForLocalizationAway @FiniteType := by |
| 33 | + introv R _ |
| 34 | + suffices Algebra.FiniteType R S by |
| 35 | + rw [RingHom.FiniteType] |
| 36 | + convert this; ext; |
| 37 | + rw [Algebra.smul_def]; rfl |
| 38 | + exact IsLocalization.finiteType_of_monoid_fg (Submonoid.powers r) S |
| 39 | +#align ring_hom.finite_type_holds_for_localization_away RingHom.finiteType_holdsForLocalizationAway |
| 40 | + |
| 41 | +theorem finiteType_ofLocalizationSpanTarget : OfLocalizationSpanTarget @FiniteType := by |
| 42 | + -- Setup algebra intances. |
| 43 | + rw [ofLocalizationSpanTarget_iff_finite] |
| 44 | + introv R hs H |
| 45 | + classical |
| 46 | + letI := f.toAlgebra |
| 47 | + replace H : ∀ r : s, Algebra.FiniteType R (Localization.Away (r : S)) |
| 48 | + · intro r; simp_rw [RingHom.FiniteType] at H; convert H r; ext; simp_rw [Algebra.smul_def]; rfl |
| 49 | + replace H := fun r => (H r).1 |
| 50 | + constructor |
| 51 | + -- Suppose `s : Finset S` spans `S`, and each `Sᵣ` is finitely generated as an `R`-algebra. |
| 52 | + -- Say `t r : Finset Sᵣ` generates `Sᵣ`. By assumption, we may find `lᵢ` such that |
| 53 | + -- `∑ lᵢ * sᵢ = 1`. I claim that all `s` and `l` and the numerators of `t` and generates `S`. |
| 54 | + choose t ht using H |
| 55 | + obtain ⟨l, hl⟩ := |
| 56 | + (Finsupp.mem_span_iff_total S (s : Set S) 1).mp |
| 57 | + (show (1 : S) ∈ Ideal.span (s : Set S) by rw [hs]; trivial) |
| 58 | + let sf := fun x : s => IsLocalization.finsetIntegerMultiple (Submonoid.powers (x : S)) (t x) |
| 59 | + use s.attach.biUnion sf ∪ s ∪ l.support.image l |
| 60 | + rw [eq_top_iff] |
| 61 | + -- We need to show that every `x` falls in the subalgebra generated by those elements. |
| 62 | + -- Since all `s` and `l` are in the subalgebra, it suffices to check that `sᵢ ^ nᵢ • x` falls in |
| 63 | + -- the algebra for each `sᵢ` and some `nᵢ`. |
| 64 | + rintro x - |
| 65 | + apply Subalgebra.mem_of_span_eq_top_of_smul_pow_mem _ (s : Set S) l hl _ _ x _ |
| 66 | + · intro x hx |
| 67 | + apply Algebra.subset_adjoin |
| 68 | + rw [Finset.coe_union, Finset.coe_union] |
| 69 | + exact Or.inl (Or.inr hx) |
| 70 | + · intro i |
| 71 | + by_cases h : l i = 0; · rw [h]; exact zero_mem _ |
| 72 | + apply Algebra.subset_adjoin |
| 73 | + rw [Finset.coe_union, Finset.coe_image] |
| 74 | + exact Or.inr (Set.mem_image_of_mem _ (Finsupp.mem_support_iff.mpr h)) |
| 75 | + · intro r |
| 76 | + rw [Finset.coe_union, Finset.coe_union, Finset.coe_biUnion] |
| 77 | + -- Since all `sᵢ` and numerators of `t r` are in the algebra, it suffices to show that the |
| 78 | + -- image of `x` in `Sᵣ` falls in the `R`-adjoin of `t r`, which is of course true. |
| 79 | + -- Porting note: The following `obtain` fails because Lean wants to know right away what the |
| 80 | + -- placeholders are, so we need to provide a little more guidance |
| 81 | + -- obtain ⟨⟨_, n₂, rfl⟩, hn₂⟩ := IsLocalization.exists_smul_mem_of_mem_adjoin |
| 82 | + -- (Submonoid.powers (r : S)) x (t r) (Algebra.adjoin R _) _ _ _ |
| 83 | + rw [show ∀ A : Set S, (∃ n, (r : S) ^ n • x ∈ Algebra.adjoin R A) ↔ |
| 84 | + (∃ m : (Submonoid.powers (r : S)), (m : S) • x ∈ Algebra.adjoin R A) by |
| 85 | + { exact fun _ => by simp [Submonoid.mem_powers_iff] }] |
| 86 | + refine IsLocalization.exists_smul_mem_of_mem_adjoin |
| 87 | + (Submonoid.powers (r : S)) x (t r) (Algebra.adjoin R _) ?_ ?_ ?_ |
| 88 | + · intro x hx |
| 89 | + apply Algebra.subset_adjoin |
| 90 | + exact Or.inl (Or.inl ⟨_, ⟨r, rfl⟩, _, ⟨s.mem_attach r, rfl⟩, hx⟩) |
| 91 | + · rw [Submonoid.powers_eq_closure, Submonoid.closure_le, Set.singleton_subset_iff] |
| 92 | + apply Algebra.subset_adjoin |
| 93 | + exact Or.inl (Or.inr r.2) |
| 94 | + · rw [ht]; trivial |
| 95 | +#align ring_hom.finite_type_of_localization_span_target RingHom.finiteType_ofLocalizationSpanTarget |
| 96 | + |
| 97 | +theorem finiteType_is_local : PropertyIsLocal @FiniteType := |
| 98 | + ⟨localization_finiteType, finiteType_ofLocalizationSpanTarget, finiteType_stableUnderComposition, |
| 99 | + finiteType_holdsForLocalizationAway⟩ |
| 100 | +#align ring_hom.finite_type_is_local RingHom.finiteType_is_local |
| 101 | + |
| 102 | +theorem finiteType_respectsIso : RingHom.RespectsIso @RingHom.FiniteType := |
| 103 | + RingHom.finiteType_is_local.respectsIso |
| 104 | +#align ring_hom.finite_type_respects_iso RingHom.finiteType_respectsIso |
| 105 | + |
| 106 | +end RingHom |
0 commit comments