Skip to content

Commit

Permalink
feat(algebraic_geometry): The underlying topological space of a Schem…
Browse files Browse the repository at this point in the history
…e is T0 (#10869)




Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com>
  • Loading branch information
erdOne and erdOne committed Dec 18, 2021
1 parent cb0a6f7 commit ac10136
Show file tree
Hide file tree
Showing 4 changed files with 89 additions and 0 deletions.
4 changes: 4 additions & 0 deletions src/algebraic_geometry/prime_spectrum/basic.lean
Expand Up @@ -657,6 +657,10 @@ lemma le_iff_mem_closure (x y : prime_spectrum R) :
by rw [← as_ideal_le_as_ideal, ← zero_locus_vanishing_ideal_eq_closure,
mem_zero_locus, vanishing_ideal_singleton, set_like.coe_subset_coe]

instance : t0_space (prime_spectrum R) :=
by { simp [t0_space_iff_or_not_mem_closure, ← le_iff_mem_closure,
← not_and_distrib, ← le_antisymm_iff, eq_comm] }

end order

end prime_spectrum
Expand Down
15 changes: 15 additions & 0 deletions src/algebraic_geometry/properties.lean
Expand Up @@ -28,6 +28,21 @@ namespace algebraic_geometry

variable (X : Scheme)

-- TODO: add sober spaces, and show that schemes are sober
instance : t0_space X.carrier :=
begin
rw t0_space_iff_distinguishable,
intros x y h h',
obtain ⟨U, R, ⟨e⟩⟩ := X.local_affine x,
have hy := (h' _ U.1.2).mp U.2,
erw ← subtype_indistinguishable_iff (⟨x, U.2⟩ : U.1.1) (⟨y, hy⟩ : U.1.1) at h',
let e' : U.1 ≃ₜ prime_spectrum R :=
homeo_of_iso ((LocallyRingedSpace.forget_to_SheafedSpace ⋙ SheafedSpace.forget _).map_iso e),
have := t0_space_of_injective_of_continuous e'.injective e'.continuous,
rw t0_space_iff_distinguishable at this,
exact this ⟨x, U.2⟩ ⟨y, hy⟩ (by simpa using h) h'
end

/-- A scheme `X` is integral if its carrier is nonempty,
and `𝒪ₓ(U)` is an integral domain for each `U ≠ ∅`. -/
class is_integral : Prop :=
Expand Down
7 changes: 7 additions & 0 deletions src/logic/basic.lean
Expand Up @@ -702,6 +702,13 @@ by rw [← decidable.not_and_distrib, decidable.not_not]

theorem and_iff_not_or_not : a ∧ b ↔ ¬ (¬ a ∨ ¬ b) := decidable.and_iff_not_or_not

@[simp] theorem not_xor (P Q : Prop) : ¬ xor P Q ↔ (P ↔ Q) :=
by simp only [not_and, xor, not_or_distrib, not_not, ← iff_iff_implies_and_implies]

theorem xor_iff_not_iff (P Q : Prop) : xor P Q ↔ ¬ (P ↔ Q) :=
by rw [iff_not_comm, not_xor]


end propositional

/-! ### Declarations about equality -/
Expand Down
63 changes: 63 additions & 0 deletions src/topology/separation.lean
Expand Up @@ -137,6 +137,41 @@ end separated
class t0_space (α : Type u) [topological_space α] : Prop :=
(t0 : ∀ x y, x ≠ y → ∃ U:set α, is_open U ∧ (xor (x ∈ U) (y ∈ U)))

lemma t0_space_def (α : Type u) [topological_space α] :
t0_space α ↔ ∀ x y, x ≠ y → ∃ U:set α, is_open U ∧ (xor (x ∈ U) (y ∈ U)) :=
by { split, apply @t0_space.t0, apply t0_space.mk }

/-- Two points are topologically indistinguishable if no open set separates them. -/
def indistinguishable {α : Type u} [topological_space α] (x y : α) : Prop :=
∀ (U : set α) (hU : is_open U), x ∈ U ↔ y ∈ U

lemma t0_space_iff_distinguishable (α : Type u) [topological_space α] :
t0_space α ↔ ∀ (x y : α), x ≠ y → ¬ indistinguishable x y :=
begin
delta indistinguishable,
rw t0_space_def,
push_neg,
simp_rw xor_iff_not_iff,
end

lemma indistinguishable_iff_closed {α : Type u} [topological_space α] (x y : α) :
indistinguishable x y ↔ ∀ (U : set α) (hU : is_closed U), x ∈ U ↔ y ∈ U :=
⟨λ h U hU, not_iff_not.mp (h _ hU.1), λ h U hU, not_iff_not.mp (h _ (is_closed_compl_iff.mpr hU))⟩

lemma indistinguishable_iff_closure {α : Type u} [topological_space α] (x y : α) :
indistinguishable x y ↔ x ∈ closure ({y} : set α) ∧ y ∈ closure ({x} : set α) :=
begin
rw indistinguishable_iff_closed,
exact ⟨λ h, ⟨(h _ is_closed_closure).mpr (subset_closure $ set.mem_singleton y),
(h _ is_closed_closure).mp (subset_closure $ set.mem_singleton x)⟩,
λ h U hU, ⟨λ hx, (is_closed.closure_subset_iff hU).mpr (set.singleton_subset_iff.mpr hx) h.2,
λ hy, (is_closed.closure_subset_iff hU).mpr (set.singleton_subset_iff.mpr hy) h.1⟩⟩
end

lemma subtype_indistinguishable_iff {α : Type u} [topological_space α] {U : set α} (x y : U) :
indistinguishable x y ↔ indistinguishable (x : α) y :=
by { simp_rw [indistinguishable_iff_closure, closure_subtype, image_singleton] }

/-- Given a closed set `S` in a compact T₀ space,
there is some `x ∈ S` such that `{x}` is closed. -/
theorem is_closed.exists_closed_singleton {α : Type*} [topological_space α]
Expand Down Expand Up @@ -207,6 +242,34 @@ instance subtype.t0_space [t0_space α] {p : α → Prop} : t0_space (subtype p)
⟨λ x y hxy, let ⟨U, hU, hxyU⟩ := t0_space.t0 (x:α) y ((not_congr subtype.ext_iff_val).1 hxy) in
⟨(coe : subtype p → α) ⁻¹' U, is_open_induced hU, hxyU⟩⟩

theorem t0_space_iff_or_not_mem_closure (α : Type u) [topological_space α] :
t0_space α ↔ (∀ a b : α, (a ≠ b) → (a ∉ closure ({b} : set α) ∨ b ∉ closure ({a} : set α))) :=
begin
simp only [← not_and_distrib, t0_space_def, not_and],
apply forall_congr, intro a,
apply forall_congr, intro b,
apply forall_congr, intro _,
split,
{ rintro ⟨s, h₁, (⟨h₂, h₃ : b ∈ sᶜ⟩|⟨h₂, h₃ : a ∈ sᶜ⟩)⟩ ha hb; rw ← is_closed_compl_iff at h₁,
{ exact (is_closed.closure_subset_iff h₁).mpr (set.singleton_subset_iff.mpr h₃) ha h₂ },
{ exact (is_closed.closure_subset_iff h₁).mpr (set.singleton_subset_iff.mpr h₃) hb h₂ } },
{ intro h,
by_cases h' : a ∈ closure ({b} : set α),
{ exact ⟨(closure {a})ᶜ, is_closed_closure.1,
or.inr ⟨h h', not_not.mpr (subset_closure (set.mem_singleton a))⟩⟩ },
{ exact ⟨(closure {b})ᶜ, is_closed_closure.1,
or.inl ⟨h', not_not.mpr (subset_closure (set.mem_singleton b))⟩⟩ } }
end

lemma t0_space_of_injective_of_continuous {α β : Type u} [topological_space α] [topological_space β]
{f : α → β} (hf : function.injective f) (hf' : continuous f) [t0_space β] : t0_space α :=
begin
constructor,
intros x y h,
obtain ⟨U, hU, e⟩ := t0_space.t0 _ _ (hf.ne h),
exact ⟨f ⁻¹' U, hf'.1 U hU, e⟩
end

/-- A T₁ space, also known as a Fréchet space, is a topological space
where every singleton set is closed. Equivalently, for every pair
`x ≠ y`, there is an open set containing `x` and not `y`. -/
Expand Down

0 comments on commit ac10136

Please sign in to comment.