From 1da7636d25b504a2155b841d1884a32c5fac1a24 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Fri, 1 Mar 2024 17:22:47 +0000 Subject: [PATCH] feat(AlgebraicGeometry/ProjectiveSpectrum/Scheme): `fromSpec` and `toSpec` compose to identity (#9618) Co-authored-by: Jeremy Tan Jie Rui --- .../ProjectiveSpectrum/Scheme.lean | 51 ++++++++++++++++++- 1 file changed, 50 insertions(+), 1 deletion(-) diff --git a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean index d2c2bda553d3f1..c68446be1b375d 100644 --- a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean +++ b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean @@ -610,7 +610,7 @@ end FromSpec section toSpecFromSpec -lemma toSpecFromSpec {f : A} {m : β„•} (hm : 0 < m) (f_deg : f ∈ π’œ m) (x : Spec.T (A⁰_ f)) : +lemma toSpec_fromSpec {f : A} {m : β„•} (hm : 0 < m) (f_deg : f ∈ π’œ m) (x : Spec.T (A⁰_ f)) : toSpec (FromSpec.toFun f_deg hm x) = x := show _ = (_ : PrimeSpectrum _) by ext (z : A⁰_ f); fconstructor <;> intro hz Β· change z ∈ ToSpec.carrier _ at hz @@ -662,6 +662,55 @@ lemma toSpecFromSpec {f : A} {m : β„•} (hm : 0 < m) (f_deg : f ∈ π’œ m) (x : end toSpecFromSpec +section fromSpecToSpec + +lemma fromSpec_toSpec {f : A} {m : β„•} (hm : 0 < m) (f_deg : f ∈ π’œ m) (x : Proj.T| pbo f) : + FromSpec.toFun f_deg hm (toSpec x) = x := by + classical + refine Subtype.ext <| ProjectiveSpectrum.ext _ _ <| HomogeneousIdeal.ext <| Ideal.ext fun z ↦ ?_ + constructor <;> intro hz + Β· rw [← DirectSum.sum_support_decompose π’œ z] + refine Ideal.sum_mem _ fun i _ ↦ ?_ + obtain ⟨c, N, acd, eq1⟩ := ToSpec.MemCarrier.clear_denominator x (hz i) + rw [HomogeneousLocalization.val_mk'', smul_mk, ← mk_one_eq_algebraMap, mk_eq_mk_iff, + r_iff_exists, OneMemClass.coe_one, one_mul] at eq1 + obtain ⟨⟨_, ⟨k, rfl⟩⟩, eq1⟩ := eq1 + dsimp at eq1 + rw [← mul_assoc, ← mul_assoc, ← pow_add, ← pow_add] at eq1 + suffices mem : f^(k + i) * βˆ‘ i in c.support.attach, acd (c i) _ * _ ∈ + x.1.asHomogeneousIdeal from + x.1.isPrime.mem_of_pow_mem _ <| x.1.isPrime.mem_or_mem (eq1.symm β–Έ mem) + |>.resolve_left fun r ↦ ProjectiveSpectrum.mem_basicOpen π’œ _ _ + |>.mp x.2 <| x.1.isPrime.mem_of_pow_mem _ r + exact Ideal.mul_mem_left _ _ <| Ideal.sum_mem _ fun i _ ↦ + Ideal.mul_mem_left _ _ i.1.2.choose_spec.1 + + Β· intro i + erw [ToSpec.mem_carrier_iff, HomogeneousLocalization.val_mk''] + dsimp only [GradedAlgebra.proj_apply] + rw [show (mk (decompose π’œ z i ^ m) ⟨f^i, ⟨i, rfl⟩⟩: Away f) = + (decompose π’œ z i ^ m : A) β€’ (mk 1 ⟨f^i, ⟨i, rfl⟩⟩ : Away f) by + Β· rw [smul_mk, smul_eq_mul, mul_one], Algebra.smul_def] + exact Ideal.mul_mem_right _ _ <| + Ideal.subset_span ⟨_, ⟨Ideal.pow_mem_of_mem _ (x.1.asHomogeneousIdeal.2 i hz) _ hm, rfl⟩⟩ + +lemma toSpec_injective {f : A} {m : β„•} (hm : 0 < m) (f_deg : f ∈ π’œ m): + Function.Injective (toSpec (π’œ := π’œ) (f := f)) := by + intro x₁ xβ‚‚ h + have := congr_arg (FromSpec.toFun f_deg hm) h + rwa [fromSpec_toSpec, fromSpec_toSpec] at this + +lemma toSpec_surjective {f : A} {m : β„•} (hm : 0 < m) (f_deg : f ∈ π’œ m): + Function.Surjective (toSpec (π’œ := π’œ) (f := f)) := + Function.surjective_iff_hasRightInverse |>.mpr + ⟨FromSpec.toFun f_deg hm, toSpec_fromSpec π’œ hm f_deg⟩ + +lemma toSpec_bijective {f : A} {m : β„•} (hm : 0 < m) (f_deg : f ∈ π’œ m): + Function.Bijective (toSpec (π’œ := π’œ) (f := f)) := + ⟨toSpec_injective π’œ hm f_deg, toSpec_surjective π’œ hm f_deg⟩ + +end fromSpecToSpec + end ProjIsoSpecTopComponent end AlgebraicGeometry