Zulip Chat Archive

Stream: maths

Topic: top space / manifold structure on points of alg varieties


Kevin Buzzard (Apr 06 2024 at 19:04):

If XX is an affine algebraic variety (or just an affine scheme of finite type) over a field KK, and if OX(X)=A\mathcal{O}_X(X)=A, and if RR is a commutative KK-algebra with a topology making RR into a topological ring, then I imagine that the RR-points X(R)X(R) (i.e. the KK-algebra homomorphisms from AA to RR) inherit a canonical topology; for example one can write down a closed immersion from XX into some large affine space AN\mathbb{A}^N and this identifies X(R)X(R) with a subspace of RNR^N, so we can give it the subspace topology. My instinct is that this topology is independent of the choice of closed immersion but I just spent a few minutes writing down what I thought would be the proof (given two embeddings into RNR^N and RMR^M consider the product embedding and then look at projections) and couldn't immediately finish it off. What trick am I missing?

Now say XX is also smooth, and KK is a nontrivially normed field. Then do the same sort of ideas give X(K)X(K) the structure of a manifold over KK, again independent of all choices? Smoothness is some criterion about differentials behaving which will be enough to make things like the implicit function theorem kick in, and then I get stuck at the same point as above.

The internet seems to be full of people saying that all of this is obvious. How hard will it be to formalise? I need it for FLT :-)

Adam Topaz (Apr 06 2024 at 20:32):

oh you don't need to embed anything. Just take the coarsest topology making the algebraic functions continuous.

Adam Topaz (Apr 06 2024 at 20:33):

you could formalize the topology as some infimum of some induced topologies, etc.

Adam Topaz (Apr 06 2024 at 20:35):

as for smoothness and the manifold structure, that's of course a much bigger deal.

Kevin Buzzard (Apr 06 2024 at 20:57):

I think that the coarsest topology making everything continuous is embedding it into affine A-space :-) So then the question becomes: is it true that this is the subspace topology for a closed embedding? Concretely, if my affine variety is GL_n and the topological ring is the reals, how do I prove that the topology you're suggesting agrees with the two topologies I've seen in the literature, namely those obtained by embedding into R^{n^2+1} via mapping M to (M,1/det(M)) and by embedding into R^{2n^2} via (M,1/M)?

Adam Topaz (Apr 06 2024 at 21:00):

but the point is that you don't even need to assume that it's affine.

Adam Topaz (Apr 06 2024 at 21:01):

If XX is any variety over KK, and UU is Zariski open, then fO(U)f \in \mathcal{O}(U) induces U(K)KU(K) \to K. Now take the induced topology with respect to this map on U(K)U(K) and take the infimum as all ff vary. Now look at the coinduced topology on X(K)X(K), then take the supermum of all such as UU varies.

Adam Topaz (Apr 06 2024 at 21:01):

I think that should work.

Adam Topaz (Apr 06 2024 at 21:05):

I think it's also possible to characterize this as the coarsest family of topologies on X(K)X(K) as XX vary over all KK-varieties for which all fO(X)f \in \mathcal{O}(X) induce continuous map X(K)KX(K) \to K and for which XYX \to Y induce continuous maps X(K)Y(K)X(K) \to Y(K) for any morphism of varieties

Adam Topaz (Apr 06 2024 at 21:07):

I guess the condition on X(K)KX(K) \to K is superfluous since you can just look at morphisms XA1X \to \mathbb{A}^1 (once you identify the topology on A1(K)=K\mathbb{A}^1(K) = K with the strong one) :)

Kevin Buzzard (Apr 06 2024 at 21:29):

Right but you're not answering the question. I don't doubt that your idea is a good idea but what I'm saying is that in the special case of GL_n(R) this is not how the topology is defined in the literature.

Kevin Buzzard (Apr 06 2024 at 21:29):

Is my guess that you can use any closed embedding (rather than all of them, which is what you're doing) not actually correct?

Kevin Buzzard (Apr 06 2024 at 21:31):

And does the same trick work for the manifold structure? I guess I can just open GAGA and see what Serre does

Adam Topaz (Apr 06 2024 at 21:41):

Oh sorry, yeah I didn't read :)

Adam Topaz (Apr 06 2024 at 21:53):

... I've been working on this #11964

Mitchell Lee (Apr 06 2024 at 22:15):

If XX is affine, I believe the topology you describe is indeed independent of the choice of closed immersion.

Let XANX \to \mathbb{A}^N be a closed immersion and let f1,,fNO(X)f_1, \ldots, f_N \in \mathcal{O}(X) be the NN coordinate functions. The topology you describe is the coarsest topology such that for all ii, the evaluation X(R)RX(R) \to R of fif_i is continuous. Since f1,,fNf_1, \ldots, f_N generate O(X)\mathcal{O}(X), this is equivalent to the statement that the evaluation X(R)RX(R) \to R of ff is continuous for all fO(X)f \in \mathcal{O}(X).

Adam Topaz (Apr 06 2024 at 22:17):

Right. You would have to show that polynomial functions are continuous for that proof. Do we have that?

Mitchell Lee (Apr 06 2024 at 22:31):

docs#MvPolynomial.continuous_eval

Mitchell Lee (Apr 06 2024 at 23:24):

In fact, I suspect that if we put a topology on X(R)X(R) using the method Adam Topaz describes, then:

  1. If XYX \to Y is a closed immersion, then X(R)Y(R)X(R) \to Y(R) is a closed embedding.
  2. (X×Y)(R)=X(R)×Y(R)(X \times Y)(R) = X(R) \times Y(R) has the product topology.

Mitchell Lee (Apr 06 2024 at 23:30):

And then, one more fact is needed to deduce the theorem you want, which is that the topology on A1(R)=R\mathbb{A}^1(R) = R is the same as the original one that you put on RR.

Adam Topaz (Apr 06 2024 at 23:40):

I’m not sure if it’s true in that level of generality. And in particular some assumptions on the ring are probably required

Adam Topaz (Apr 06 2024 at 23:44):

Cf. https://arxiv.org/pdf/2009.02319.pdf

Adam Topaz (Apr 06 2024 at 23:49):

Well Fact 4.7 shows that it’s ok in the case of a field

Adam Topaz (Apr 06 2024 at 23:49):

(Which comes from the red book)

Adam Topaz (Apr 07 2024 at 00:01):

Ok, now I’m curious to see in what level of generality this holds :)

Adam Topaz (Apr 07 2024 at 00:03):

What if we put a topology on O(S)\mathcal{O}(S) for some scheme S. will it hold for the X(S)’s for X separated and of finite type over S?

Kevin Buzzard (Apr 07 2024 at 00:32):

I went for topological rings because I wanted to apply this for the finite adeles of a number field, which are hideously non-Noetherian.

Adam Topaz (Apr 07 2024 at 01:24):

@Kevin Buzzard I found these notes by Conrad which makes things pretty explicit: https://math.stanford.edu/~conrad/papers/adelictop.pdf

Kevin Buzzard (Apr 07 2024 at 11:06):

oh many thanks! I had considered asking Brian instead of asking here, but I've been badgering him about finite flat group schemes for a while now so thought I'd give him a break :-)

Antoine Chambert-Loir (Apr 07 2024 at 13:46):

I was going to give the same reference. It is not so trivial, though, as the case of adeles can suggest.

Mitchell Lee (Apr 07 2024 at 20:13):

For the manifold question, here is the "low-powered" way. Suppose you have two closed immersions XAmX \to \mathbb{A}^m and XAnX \to \mathbb{A}^n. We wish to show that the corresponding closed immersions f ⁣:X(K)Kmf \colon X(K) \to K^m and g ⁣:X(K)Kng \colon X(K) \to K^n induce the same analytic manifold structure on X(K)X(K).

For any xX(K)x \in X(K), one can construct an analytic function h ⁣:UKnh \colon U \to K^n, defined on a neighborhood UU of f(x)f(x) in KmK^m, such that h(f(y))=g(y)h(f(y)) = g(y) for yy near xx. (One can construct hh by composing gg with a function that is locally a left inverse of ff. This exists by the implicit function theorem.) So the identity function idX(K)\operatorname{id}_{X(K)} is locally analytic at xx, where the manifold structure on the domain is induced by ff and the manifold structure on the codomain is induced by gg. You can do the same in the reverse direction.

Mitchell Lee (Apr 07 2024 at 20:24):

Hopefully I have not misunderstood what is going on here.

Kevin Buzzard (Apr 07 2024 at 20:24):

This makes me think that there should be an analogous low-level argument to show that any two closed immersions X -> A^n should induce the same topology on X(R).

Mitchell Lee (Apr 07 2024 at 20:32):

Now that I think about this argument, maybe it needs a little more work.

I am referring to the work required to prove that hh is actually analytic. I think you can't just say "it's analytic because it's the composition of two analytic functions", because that would be circular reasoning.

Mitchell Lee (Apr 07 2024 at 20:39):

Even if this works, I'm not convinced there's an analogous argument for the topology on X(R)X(R). This seems to intrinsically rely on the implicit function theorem.

Antoine Chambert-Loir (Apr 08 2024 at 08:34):

Kevin Buzzard said:

This makes me think that there should be an analogous low-level argument to show that any two closed immersions X -> A^n should induce the same topology on X(R).

Take the diagonal embedding in A^m x A^n. Taking the induced topologies of the product topology of R gives you 3 topologies on X, say T1 T2 and T. The maps from (X,T) to (X,T1) and (X,T2) are continuous since they just forget about coordinates, and one needs to show that they are homeomorphisms. Now, by definition of a closed immersion (Nullstellensatz), the second embedding of X to A^n factors through a map A^m -> A^n. Consequently, (X,T1)->(X,T) is of the form (x1,...,xm) -> (x1,...,xm, f1(x1,..,xm),...,fn(x1,...,xm)) where f1,..., fn are polynomials, and it follows from the definition of a topological ring that they induce continuous functions R^m ->R.

Mitchell Lee (Apr 08 2024 at 16:39):

I didn't realize that you could just factor the map algebraically. That also gives a cleaner proof that the analytic structure is preserved.


Last updated: May 02 2025 at 03:31 UTC