Skip to content

Commit

Permalink
This commit does not belong to any branch on this repository, and may belong to a fork outside of the repository.
make directory
Browse files Browse the repository at this point in the history
mzeuner committed Feb 6, 2024
1 parent bf3debd commit 9741a31
Showing 11 changed files with 978 additions and 835 deletions.
239 changes: 239 additions & 0 deletions Cubical/AlgebraicGeometry/Functorial/ZFunctors/Base.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,239 @@
{-
A ℤ-functor is just a functor from rings to sets.
NOTE: we consider the functor category [ Ring ℓ , Set ℓ ] for some universe level ℓ
and not [ Ring ℓ , Set (ℓ+1) ] as is done in
"Introduction to Algebraic Geometry and Algebraic Groups"
by Demazure & Gabriel!
The category of ℤ-functors contains the category of (qcqs-) schemes
as a full subcategory and satisfies a "universal property"
similar to the one of schemes:
There is an adjunction 𝓞 ⊣ᵢ Sp
(relative to the inclusion i : CommRing ℓ → CommRing (ℓ+1))
between the "global sections functor" 𝓞
and the fully-faithful embedding of affines Sp,
whose counit is a natural isomorphism
-}


{-# OPTIONS --safe --lossy-unification #-}
module Cubical.AlgebraicGeometry.Functorial.ZFunctors.Base where

open import Cubical.Foundations.Prelude
open import Cubical.Foundations.Equiv
open import Cubical.Foundations.Isomorphism
open import Cubical.Foundations.Function
open import Cubical.Foundations.Powerset
open import Cubical.Foundations.HLevels

open import Cubical.Functions.FunExtEquiv

open import Cubical.Data.Sigma
open import Cubical.Data.Nat using (ℕ)

open import Cubical.Algebra.Ring
open import Cubical.Algebra.CommRing

open import Cubical.Categories.Category renaming (isIso to isIsoC)
open import Cubical.Categories.Functor
open import Cubical.Categories.Instances.Sets
open import Cubical.Categories.Instances.CommRings
open import Cubical.Categories.Instances.Functors
open import Cubical.Categories.NaturalTransformation
open import Cubical.Categories.Yoneda
open import Cubical.Categories.Site.Sheaf
open import Cubical.Categories.Site.Instances.ZariskiCommRing

open import Cubical.HITs.PropositionalTruncation as PT


open Category hiding (_∘_) renaming (_⋆_ to _⋆c_)


module _ {ℓ : Level} where

open Functor
open NatTrans
open CommRingStr ⦃...⦄
open IsRingHom

-- using the naming conventions of Demazure & Gabriel
ℤFunctor = Functor (CommRingsCategory {ℓ = ℓ}) (SET ℓ)
ℤFUNCTOR = FUNCTOR (CommRingsCategory {ℓ = ℓ}) (SET ℓ)

-- uses that double op is original category definitionally
Sp : Functor (CommRingsCategory {ℓ = ℓ} ^op) ℤFUNCTOR
Sp = YO {C = (CommRingsCategory {ℓ = ℓ} ^op)}

-- Affine schemes are merely representable functors
-- TODO: by univalence it should be fine to take untruncated Σ
isAffine : ℤFunctor Type (ℓ-suc ℓ)
isAffine X = ∃[ A ∈ CommRing ℓ ] NatIso (Sp .F-ob A) X

-- a ℤ-functor that is a sheaf wrt the Zariski coverage is called local
isLocal : ℤFunctor Type (ℓ-suc ℓ)
isLocal X = isSheaf zariskiCoverage X

-- the forgetful functor
-- aka the affine line
-- (aka the representable of ℤ[x])
𝔸¹ : ℤFunctor
𝔸¹ = ForgetfulCommRing→Set

-- the "global sections" functor
𝓞 : Functor ℤFUNCTOR (CommRingsCategory {ℓ = ℓ-suc ℓ} ^op)
fst (F-ob 𝓞 X) = X ⇒ 𝔸¹

-- ring struncture induced by internal ring object 𝔸¹
N-ob (CommRingStr.0r (snd (F-ob 𝓞 X))) A _ = 0r
where instance _ = A .snd
N-hom (CommRingStr.0r (snd (F-ob 𝓞 X))) φ = funExt λ _ sym (φ .snd .pres0)

N-ob (CommRingStr.1r (snd (F-ob 𝓞 X))) A _ = 1r
where instance _ = A .snd
N-hom (CommRingStr.1r (snd (F-ob 𝓞 X))) φ = funExt λ _ sym (φ .snd .pres1)

N-ob ((snd (F-ob 𝓞 X) CommRingStr.+ α) β) A x = α .N-ob A x + β .N-ob A x
where instance _ = A .snd
N-hom ((snd (F-ob 𝓞 X) CommRingStr.+ α) β) {x = A} {y = B} φ = funExt path
where
instance
_ = A .snd
_ = B .snd
path : x α .N-ob B (X .F-hom φ x) + β .N-ob B (X .F-hom φ x)
≡ φ .fst (α .N-ob A x + β .N-ob A x)
path x = α .N-ob B (X .F-hom φ x) + β .N-ob B (X .F-hom φ x)
≡⟨ cong₂ _+_ (funExt⁻ (α .N-hom φ) x) (funExt⁻ (β .N-hom φ) x) ⟩
φ .fst (α .N-ob A x) + φ .fst (β .N-ob A x)
≡⟨ sym (φ .snd .pres+ _ _) ⟩
φ .fst (α .N-ob A x + β .N-ob A x) ∎

N-ob ((snd (F-ob 𝓞 X) CommRingStr.· α) β) A x = α .N-ob A x · β .N-ob A x
where instance _ = A .snd
N-hom ((snd (F-ob 𝓞 X) CommRingStr.· α) β) {x = A} {y = B} φ = funExt path
where
instance
_ = A .snd
_ = B .snd
path : x α .N-ob B (X .F-hom φ x) · β .N-ob B (X .F-hom φ x)
≡ φ .fst (α .N-ob A x · β .N-ob A x)
path x = α .N-ob B (X .F-hom φ x) · β .N-ob B (X .F-hom φ x)
≡⟨ cong₂ _·_ (funExt⁻ (α .N-hom φ) x) (funExt⁻ (β .N-hom φ) x) ⟩
φ .fst (α .N-ob A x) · φ .fst (β .N-ob A x)
≡⟨ sym (φ .snd .pres· _ _) ⟩
φ .fst (α .N-ob A x · β .N-ob A x) ∎

N-ob ((CommRingStr.- snd (F-ob 𝓞 X)) α) A x = - α .N-ob A x
where instance _ = A .snd
N-hom ((CommRingStr.- snd (F-ob 𝓞 X)) α) {x = A} {y = B} φ = funExt path
where
instance
_ = A .snd
_ = B .snd
path : x - α .N-ob B (X .F-hom φ x) ≡ φ .fst (- α .N-ob A x)
path x = - α .N-ob B (X .F-hom φ x) ≡⟨ cong -_ (funExt⁻ (α .N-hom φ) x) ⟩
- φ .fst (α .N-ob A x) ≡⟨ sym (φ .snd .pres- _) ⟩
φ .fst (- α .N-ob A x) ∎

CommRingStr.isCommRing (snd (F-ob 𝓞 X)) = makeIsCommRing
isSetNatTrans
(λ _ _ _ makeNatTransPath (funExt₂ λ A _ A .snd .CommRingStr.+Assoc _ _ _))
(λ _ makeNatTransPath (funExt₂ λ A _ A .snd .CommRingStr.+IdR _))
(λ _ makeNatTransPath (funExt₂ λ A _ A .snd .CommRingStr.+InvR _))
(λ _ _ makeNatTransPath (funExt₂ λ A _ A .snd .CommRingStr.+Comm _ _))
(λ _ _ _ makeNatTransPath (funExt₂ λ A _ A .snd .CommRingStr.·Assoc _ _ _))
(λ _ makeNatTransPath (funExt₂ λ A _ A .snd .CommRingStr.·IdR _))
(λ _ _ _ makeNatTransPath (funExt₂ λ A _ A .snd .CommRingStr.·DistR+ _ _ _))
(λ _ _ makeNatTransPath (funExt₂ λ A _ A .snd .CommRingStr.·Comm _ _))

-- action on natural transformations
fst (F-hom 𝓞 α) = α ●ᵛ_
pres0 (snd (F-hom 𝓞 α)) = makeNatTransPath (funExt₂ λ _ _ refl)
pres1 (snd (F-hom 𝓞 α)) = makeNatTransPath (funExt₂ λ _ _ refl)
pres+ (snd (F-hom 𝓞 α)) _ _ = makeNatTransPath (funExt₂ λ _ _ refl)
pres· (snd (F-hom 𝓞 α)) _ _ = makeNatTransPath (funExt₂ λ _ _ refl)
pres- (snd (F-hom 𝓞 α)) _ = makeNatTransPath (funExt₂ λ _ _ refl)

-- functoriality of 𝓞
F-id 𝓞 = RingHom≡ (funExt λ _ makeNatTransPath (funExt₂ λ _ _ refl))
F-seq 𝓞 _ _ = RingHom≡ (funExt λ _ makeNatTransPath (funExt₂ λ _ _ refl))


-- we get a relative adjunction 𝓞 ⊣ᵢ Sp
-- with respect to the inclusion i : CommRing ℓ → CommRing (ℓ+1)
module AdjBij {ℓ : Level} where

open Functor
open NatTrans
open Iso
open IsRingHom

private module _ {A : CommRing ℓ} {X : ℤFunctor} where
_♭ : CommRingHom A (𝓞 .F-ob X) X ⇒ Sp .F-ob A
fst (N-ob (φ ♭) B x) a = φ .fst a .N-ob B x

pres0 (snd (N-ob (φ ♭) B x)) = cong (λ y y .N-ob B x) (φ .snd .pres0)
pres1 (snd (N-ob (φ ♭) B x)) = cong (λ y y .N-ob B x) (φ .snd .pres1)
pres+ (snd (N-ob (φ ♭) B x)) _ _ = cong (λ y y .N-ob B x) (φ .snd .pres+ _ _)
pres· (snd (N-ob (φ ♭) B x)) _ _ = cong (λ y y .N-ob B x) (φ .snd .pres· _ _)
pres- (snd (N-ob (φ ♭) B x)) _ = cong (λ y y .N-ob B x) (φ .snd .pres- _)

N-hom (φ ♭) ψ = funExt (λ x RingHom≡ (funExt λ a funExt⁻ (φ .fst a .N-hom ψ) x))


-- the other direction is just precomposition modulo Yoneda
_♯ : X ⇒ Sp .F-ob A CommRingHom A (𝓞 .F-ob X)
fst (α ♯) a = α ●ᵛ yonedaᴾ 𝔸¹ A .inv a

pres0 (snd (α ♯)) = makeNatTransPath (funExt₂ λ B x α .N-ob B x .snd .pres0)
pres1 (snd (α ♯)) = makeNatTransPath (funExt₂ λ B x α .N-ob B x .snd .pres1)
pres+ (snd (α ♯)) _ _ = makeNatTransPath (funExt₂ λ B x α .N-ob B x .snd .pres+ _ _)
pres· (snd (α ♯)) _ _ = makeNatTransPath (funExt₂ λ B x α .N-ob B x .snd .pres· _ _)
pres- (snd (α ♯)) _ = makeNatTransPath (funExt₂ λ B x α .N-ob B x .snd .pres- _)


-- the two maps are inverse to each other
♭♯Id : : X ⇒ Sp .F-ob A) ((α ♯) ♭) ≡ α
♭♯Id _ = makeNatTransPath (funExt₂ λ _ _ RingHom≡ (funExt (λ _ refl)))

♯♭Id : : CommRingHom A (𝓞 .F-ob X)) ((φ ♭) ♯) ≡ φ
♯♭Id _ = RingHom≡ (funExt λ _ makeNatTransPath (funExt₂ λ _ _ refl))

𝓞⊣SpIso : {A : CommRing ℓ} {X : ℤFunctor {ℓ}}
Iso (CommRingHom A (𝓞 .F-ob X)) (X ⇒ Sp .F-ob A)
fun 𝓞⊣SpIso = _♭
inv 𝓞⊣SpIso = _♯
rightInv 𝓞⊣SpIso = ♭♯Id
leftInv 𝓞⊣SpIso = ♯♭Id

𝓞⊣SpNatℤFunctor : {A : CommRing ℓ} {X Y : ℤFunctor {ℓ}} (α : X ⇒ Sp .F-ob A) (β : Y ⇒ X)
(β ●ᵛ α) ♯ ≡ (𝓞 .F-hom β) ∘cr (α ♯)
𝓞⊣SpNatℤFunctor _ _ = RingHom≡ (funExt (λ _ makeNatTransPath (funExt₂ (λ _ _ refl))))

𝓞⊣SpNatCommRing : {X : ℤFunctor {ℓ}} {A B : CommRing ℓ}
: CommRingHom A (𝓞 .F-ob X)) (ψ : CommRingHom B A)
(φ ∘cr ψ) ♭ ≡ (φ ♭) ●ᵛ Sp .F-hom ψ
𝓞⊣SpNatCommRing _ _ = makeNatTransPath (funExt₂ λ _ _ RingHom≡ (funExt (λ _ refl)))

-- right relative adjunction induces a counit,
-- which is an equivalence
private
ε : (A : CommRing ℓ) CommRingHom A ((𝓞 ∘F Sp) .F-ob A)
ε A = (idTrans (Sp .F-ob A)) ♯

𝓞⊣SpCounitEquiv : (A : CommRing ℓ) CommRingEquiv A ((𝓞 ∘F Sp) .F-ob A)
fst (𝓞⊣SpCounitEquiv A) = isoToEquiv theIso
where
theIso : Iso (A .fst) ((𝓞 ∘F Sp) .F-ob A .fst)
fun theIso = ε A .fst
inv theIso = yonedaᴾ 𝔸¹ A .fun
rightInv theIso α = ℤFUNCTOR .⋆IdL _ ∙ yonedaᴾ 𝔸¹ A .leftInv α
leftInv theIso a = path -- I get yellow otherwise
where
path : yonedaᴾ 𝔸¹ A .fun ((idTrans (Sp .F-ob A)) ●ᵛ yonedaᴾ 𝔸¹ A .inv a) ≡ a
path = cong (yonedaᴾ 𝔸¹ A .fun) (ℤFUNCTOR .⋆IdL _) ∙ yonedaᴾ 𝔸¹ A .rightInv a
snd (𝓞⊣SpCounitEquiv A) = ε A .snd
394 changes: 394 additions & 0 deletions Cubical/AlgebraicGeometry/Functorial/ZFunctors/CompactOpen.agda

Large diffs are not rendered by default.

208 changes: 208 additions & 0 deletions Cubical/AlgebraicGeometry/Functorial/ZFunctors/OpenSubscheme.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,208 @@
{-
Compact open subfunctors of qcqs-schemes are qcqs-schemes (TODO!!!)
The proof proceeds by
1. Defining standard/basic compact opens of affines and proving that they are affine
2. Proving that arbitrary compact opens of affines a re qcqs-schemes
-}
{-# OPTIONS --safe --lossy-unification #-}
module Cubical.AlgebraicGeometry.Functorial.ZFunctors.OpenSubscheme where

open import Cubical.Foundations.Prelude
open import Cubical.Foundations.Equiv
open import Cubical.Foundations.Isomorphism
open import Cubical.Foundations.Function
open import Cubical.Foundations.Powerset
open import Cubical.Foundations.HLevels


open import Cubical.Functions.FunExtEquiv

open import Cubical.Data.Sigma
open import Cubical.Data.Nat using (ℕ)

open import Cubical.Data.FinData

open import Cubical.Algebra.Ring
open import Cubical.Algebra.CommRing
open import Cubical.Algebra.CommRing.Localisation
open import Cubical.Algebra.CommRing.RadicalIdeal
open import Cubical.Algebra.Semilattice
open import Cubical.Algebra.Lattice
open import Cubical.Algebra.DistLattice
open import Cubical.Algebra.DistLattice.BigOps

open import Cubical.AlgebraicGeometry.ZariskiLattice.Base
open import Cubical.AlgebraicGeometry.ZariskiLattice.UniversalProperty
open import Cubical.AlgebraicGeometry.ZariskiLattice.Properties
open import Cubical.AlgebraicGeometry.Functorial.ZFunctors.Base
open import Cubical.AlgebraicGeometry.Functorial.ZFunctors.CompactOpen
open import Cubical.AlgebraicGeometry.Functorial.ZFunctors.QcQsScheme

open import Cubical.Categories.Category renaming (isIso to isIsoC)
open import Cubical.Categories.Functor
open import Cubical.Categories.Instances.Sets
open import Cubical.Categories.Instances.CommRings
open import Cubical.Categories.Instances.DistLattice
open import Cubical.Categories.Instances.DistLattices
open import Cubical.Categories.Instances.Functors
open import Cubical.Categories.Site.Cover
open import Cubical.Categories.Site.Coverage
open import Cubical.Categories.Site.Sheaf
open import Cubical.Categories.Site.Instances.ZariskiCommRing
open import Cubical.Categories.NaturalTransformation
open import Cubical.Categories.Yoneda


open import Cubical.HITs.PropositionalTruncation as PT
open import Cubical.HITs.SetQuotients as SQ

open import Cubical.Relation.Binary.Order.Poset

open Category hiding (_∘_) renaming (_⋆_ to _⋆c_)



-- standard affine opens
module _ {ℓ : Level} (R : CommRing ℓ) (f : R .fst) where

open Iso
open Functor
open NatTrans
open NatIso
open isIsoC
open DistLatticeStr ⦃...⦄
open CommRingStr ⦃...⦄
open IsRingHom
open RingHoms
open IsLatticeHom
open ZarLat

open InvertingElementsBase R
open UniversalProp f

private module ZL = ZarLatUniversalProp

private instance _ = R .snd

D : CompactOpen (Sp ⟅ R ⟆)
D = yonedaᴾ ZarLatFun R .inv (ZL.D R f)

SpR[1/f]≅⟦Df⟧ : NatIso (Sp .F-ob R[1/ f ]AsCommRing) ⟦ D ⟧ᶜᵒ
N-ob (trans SpR[1/f]≅⟦Df⟧) B φ = (φ ∘r /1AsCommRingHom) , ∨lRid _ ∙ path
where
open CommRingHomTheory φ
open IsSupport (ZL.isSupportD B)
instance
_ = B .snd
_ = ZariskiLattice B .snd

isUnitφ[f/1] : φ .fst (f /1) ∈ B ˣ
isUnitφ[f/1] = RingHomRespInv (f /1) ⦃ S/1⊆S⁻¹Rˣ f ∣ 1 , sym (·IdR f) ∣₁ ⦄

path : ZL.D B (φ .fst (f /1)) ≡ 1l
path = supportUnit _ isUnitφ[f/1]

N-hom (trans SpR[1/f]≅⟦Df⟧) _ = funExt λ _ Σ≡Prop (λ _ squash/ _ _) (RingHom≡ refl)

inv (nIso SpR[1/f]≅⟦Df⟧ B) (φ , Dφf≡D1) = invElemUniversalProp B φ isUnitφf .fst .fst
where
instance _ = ZariskiLattice B .snd
isUnitφf : φ .fst f ∈ B ˣ
isUnitφf = unitLemmaZarLat B (φ $r f) (sym (∨lRid _) ∙ Dφf≡D1)

sec (nIso SpR[1/f]≅⟦Df⟧ B) =
funExt λ _ Σ≡Prop (λ _ squash/ _ _) (RingHom≡ (invElemUniversalProp _ _ _ .fst .snd))
ret (nIso SpR[1/f]≅⟦Df⟧ B) =
funExt λ φ cong fst (invElemUniversalProp B (φ ∘r /1AsCommRingHom) _ .snd (φ , refl))

isAffineD : isAffineCompactOpen D
isAffineD = ∣ R[1/ f ]AsCommRing , SpR[1/f]≅⟦Df⟧ ∣₁


-- compact opens of affine schemes are qcqs-schemes
module _ {ℓ : Level} (R : CommRing ℓ) (W : CompactOpen (Sp ⟅ R ⟆)) where

open Iso
open Functor
open NatTrans
open NatIso
open isIsoC
open DistLatticeStr ⦃...⦄
open CommRingStr ⦃...⦄
open PosetStr ⦃...⦄
open IsRingHom
open RingHoms
open IsLatticeHom
open ZarLat


open JoinSemilattice (Lattice→JoinSemilattice (DistLattice→Lattice (CompOpenDistLattice .F-ob (Sp .F-ob R)))) using (IndPoset; ind≤bigOp)
open InvertingElementsBase R
open Join
open JoinMap
open AffineCover
module ZL = ZarLatUniversalProp

private
instance
_ = R .snd
_ = ZariskiLattice R .snd
_ = CompOpenDistLattice .F-ob (Sp .F-ob R) .snd
_ = CompOpenDistLattice .F-ob ⟦ W ⟧ᶜᵒ .snd
_ = IndPoset .snd

w : ZL R
w = yonedaᴾ ZarLatFun R .fun W

-- yoneda is a lattice homomorphsim
isHomYoneda : IsLatticeHom (DistLattice→Lattice (ZariskiLattice R) .snd)
(yonedaᴾ ZarLatFun R .inv)
(DistLattice→Lattice (CompOpenDistLattice ⟅ Sp ⟅ R ⟆ ⟆) .snd)
pres0 isHomYoneda = makeNatTransPath (funExt₂ (λ _ _ refl))
pres1 isHomYoneda =
makeNatTransPath (funExt₂ (λ _ φ inducedZarLatHom φ .snd .pres1))
pres∨l isHomYoneda u v =
makeNatTransPath (funExt₂ (λ _ φ inducedZarLatHom φ .snd .pres∨l u v))
pres∧l isHomYoneda u v =
makeNatTransPath (funExt₂ (λ _ φ inducedZarLatHom φ .snd .pres∧l u v))

module _ {n : ℕ}
(f : FinVec (fst R) n)
(⋁Df≡W : ⋁ (CompOpenDistLattice ⟅ Sp ⟅ R ⟆ ⟆) (D R ∘ f) ≡ W) where

Df≤W : i D R (f i) ≤ W
Df≤W i = subst (D R (f i) ≤_) ⋁Df≡W (ind≤bigOp (D R ∘ f) i)

toAffineCover : AffineCover ⟦ W ⟧ᶜᵒ
AffineCover.n toAffineCover = n
U toAffineCover i = compOpenDownHom (Sp ⟅ R ⟆) W .fst (D R (f i))
covers toAffineCover = sym (pres⋁ (compOpenDownHom (Sp ⟅ R ⟆) W) (D R ∘ f))
∙ cong (compOpenDownHom (Sp ⟅ R ⟆) W .fst) ⋁Df≡W
∙ makeNatTransPath (funExt₂ (λ _ snd))
isAffineU toAffineCover i =
∣ _ , seqNatIso (SpR[1/f]≅⟦Df⟧ R (f i)) (compOpenDownHomNatIso _ (Df≤W i)) ∣₁

module _ {n : ℕ}
(f : FinVec (fst R) n)
(⋁Df≡w : ⋁ (ZariskiLattice R) (ZL.D R ∘ f) ≡ w) where

private
⋁Df≡W : ⋁ (CompOpenDistLattice ⟅ Sp ⟅ R ⟆ ⟆) (D R ∘ f) ≡ W
⋁Df≡W = sym (pres⋁ (_ , isHomYoneda) (ZL.D R ∘ f))
∙ cong (yonedaᴾ ZarLatFun R .inv) ⋁Df≡w
∙ yonedaᴾ ZarLatFun R .leftInv W

makeAffineCoverCompOpenOfAffine : AffineCover ⟦ W ⟧ᶜᵒ
makeAffineCoverCompOpenOfAffine = toAffineCover f ⋁Df≡W

hasAffineCoverCompOpenOfAffine : hasAffineCover ⟦ W ⟧ᶜᵒ
hasAffineCoverCompOpenOfAffine = PT.map truncHelper ([]surjective w)
where
truncHelper : Σ[ n,f ∈ Σ ℕ (FinVec (fst R)) ] [ n,f ] ≡ w AffineCover ⟦ W ⟧ᶜᵒ
truncHelper ((n , f) , [n,f]≡w) = makeAffineCoverCompOpenOfAffine f (ZL.⋁D≡ R f ∙ [n,f]≡w)

isQcQsSchemeCompOpenOfAffine : isQcQsScheme ⟦ W ⟧ᶜᵒ
fst isQcQsSchemeCompOpenOfAffine = presLocalCompactOpen _ _ (isSubcanonicalZariskiCoverage R)
snd isQcQsSchemeCompOpenOfAffine = hasAffineCoverCompOpenOfAffine
97 changes: 97 additions & 0 deletions Cubical/AlgebraicGeometry/Functorial/ZFunctors/QcQsScheme.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,97 @@
{-
A qcqs-scheme is a ℤ-functor that is a local (a Zariski-sheaf)
and has an affine cover, where that notion of cover is given
by the lattice structure of compact open subfunctors
-}

{-# OPTIONS --safe --lossy-unification #-}
module Cubical.AlgebraicGeometry.Functorial.ZFunctors.QcQsScheme where

open import Cubical.Foundations.Prelude
open import Cubical.Foundations.Equiv
open import Cubical.Foundations.Isomorphism
open import Cubical.Foundations.Function
open import Cubical.Foundations.Powerset
open import Cubical.Foundations.HLevels


open import Cubical.Functions.FunExtEquiv

open import Cubical.Data.Sigma
open import Cubical.Data.Nat using (ℕ)

open import Cubical.Data.FinData

open import Cubical.Algebra.Ring
open import Cubical.Algebra.CommRing
open import Cubical.Algebra.CommRing.Localisation
open import Cubical.Algebra.CommRing.RadicalIdeal
open import Cubical.Algebra.Semilattice
open import Cubical.Algebra.Lattice
open import Cubical.Algebra.DistLattice
open import Cubical.Algebra.DistLattice.BigOps

open import Cubical.AlgebraicGeometry.Functorial.ZFunctors.Base
open import Cubical.AlgebraicGeometry.Functorial.ZFunctors.CompactOpen

open import Cubical.Categories.Category renaming (isIso to isIsoC)
open import Cubical.Categories.Functor
open import Cubical.Categories.Instances.Sets
open import Cubical.Categories.Instances.CommRings
open import Cubical.Categories.Instances.DistLattice
open import Cubical.Categories.Instances.DistLattices
open import Cubical.Categories.Instances.Functors
open import Cubical.Categories.Site.Cover
open import Cubical.Categories.Site.Coverage
open import Cubical.Categories.Site.Sheaf
open import Cubical.Categories.Site.Instances.ZariskiCommRing
open import Cubical.Categories.NaturalTransformation
open import Cubical.Categories.Yoneda


open import Cubical.HITs.PropositionalTruncation as PT

open import Cubical.Relation.Binary.Order.Poset

open Category hiding (_∘_) renaming (_⋆_ to _⋆c_)

module _ {ℓ : Level} (X : ℤFunctor {ℓ = ℓ}) where
open Functor
open DistLatticeStr ⦃...⦄
open Join (CompOpenDistLattice .F-ob X)
open JoinSemilattice (Lattice→JoinSemilattice (DistLattice→Lattice (CompOpenDistLattice .F-ob X)))
private instance _ = (CompOpenDistLattice .F-ob X) .snd

record AffineCover : Type (ℓ-suc ℓ) where
field
n :
U : FinVec (CompactOpen X) n
covers : ⋁ U ≡ 1l -- TODO: equivalent to X ≡ ⟦ ⋁ U ⟧ᶜᵒ
isAffineU : i isAffineCompactOpen (U i)

hasAffineCover : Type (ℓ-suc ℓ)
hasAffineCover = ∥ AffineCover ∥₁

isQcQsScheme : Type (ℓ-suc ℓ)
isQcQsScheme = isLocal X × hasAffineCover


-- affine schemes are qcqs-schemes
module _ {ℓ : Level} (A : CommRing ℓ) where
open Functor
open DistLatticeStr ⦃...⦄
open AffineCover
private instance _ = (CompOpenDistLattice .F-ob (Sp .F-ob A)) .snd

-- the canonical one element affine cover of a representable
singlAffineCover : AffineCover (Sp .F-ob A)
n singlAffineCover = 1
U singlAffineCover zero = 1l
covers singlAffineCover = ∨lRid _
isAffineU singlAffineCover zero = ∣ A , X≅⟦1⟧ (Sp .F-ob A) ∣₁

isQcQsSchemeAffine : isQcQsScheme (Sp .F-ob A)
fst isQcQsSchemeAffine = isSubcanonicalZariskiCoverage A
snd isQcQsSchemeAffine = ∣ singlAffineCover ∣₁
Original file line number Diff line number Diff line change
@@ -1,5 +1,5 @@
{-# OPTIONS --safe --lossy-unification #-}
module Cubical.Algebra.ZariskiLattice.Base where
module Cubical.AlgebraicGeometry.ZariskiLattice.Base where


open import Cubical.Foundations.Prelude
Original file line number Diff line number Diff line change
@@ -1,5 +1,5 @@
{-# OPTIONS --safe --lossy-unification #-}
module Cubical.Algebra.ZariskiLattice.Properties where
module Cubical.AlgebraicGeometry.ZariskiLattice.Properties where


open import Cubical.Foundations.Prelude
@@ -44,8 +44,8 @@ open import Cubical.Algebra.DistLattice.BigOps
open import Cubical.Algebra.DistLattice.Downset
open import Cubical.Algebra.Matrix

open import Cubical.Algebra.ZariskiLattice.Base
open import Cubical.Algebra.ZariskiLattice.UniversalProperty
open import Cubical.AlgebraicGeometry.ZariskiLattice.Base
open import Cubical.AlgebraicGeometry.ZariskiLattice.UniversalProperty

open import Cubical.HITs.SetQuotients as SQ
open import Cubical.HITs.PropositionalTruncation as PT
Original file line number Diff line number Diff line change
@@ -8,7 +8,7 @@
-}

{-# OPTIONS --safe --lossy-unification #-}
module Cubical.Algebra.ZariskiLattice.StructureSheaf where
module Cubical.AlgebraicGeometry.ZariskiLattice.StructureSheaf where


open import Cubical.Foundations.Prelude
@@ -56,8 +56,9 @@ open import Cubical.Algebra.Lattice
open import Cubical.Algebra.DistLattice
open import Cubical.Algebra.DistLattice.Basis
open import Cubical.Algebra.DistLattice.BigOps
open import Cubical.Algebra.ZariskiLattice.Base
open import Cubical.Algebra.ZariskiLattice.UniversalProperty

open import Cubical.AlgebraicGeometry.ZariskiLattice.Base
open import Cubical.AlgebraicGeometry.ZariskiLattice.UniversalProperty

open import Cubical.Categories.Category.Base hiding (_[_,_])
open import Cubical.Categories.Functor
Original file line number Diff line number Diff line change
@@ -9,7 +9,7 @@
-}

{-# OPTIONS --safe --lossy-unification #-}
module Cubical.Algebra.ZariskiLattice.StructureSheafPullback where
module Cubical.AlgebraicGeometry.ZariskiLattice.StructureSheafPullback where


open import Cubical.Foundations.Prelude
@@ -59,8 +59,9 @@ open import Cubical.Algebra.Lattice
open import Cubical.Algebra.DistLattice
open import Cubical.Algebra.DistLattice.Basis
open import Cubical.Algebra.DistLattice.BigOps
open import Cubical.Algebra.ZariskiLattice.Base
open import Cubical.Algebra.ZariskiLattice.UniversalProperty

open import Cubical.AlgebraicGeometry.ZariskiLattice.Base
open import Cubical.AlgebraicGeometry.ZariskiLattice.UniversalProperty

open import Cubical.Categories.Category.Base hiding (_[_,_])
open import Cubical.Categories.Functor
Original file line number Diff line number Diff line change
@@ -1,5 +1,5 @@
{-# OPTIONS --safe --lossy-unification #-}
module Cubical.Algebra.ZariskiLattice.UniversalProperty where
module Cubical.AlgebraicGeometry.ZariskiLattice.UniversalProperty where


open import Cubical.Foundations.Prelude
@@ -41,7 +41,7 @@ open import Cubical.Algebra.DistLattice.Basis
open import Cubical.Algebra.DistLattice.BigOps
open import Cubical.Algebra.Matrix

open import Cubical.Algebra.ZariskiLattice.Base
open import Cubical.AlgebraicGeometry.ZariskiLattice.Base

open import Cubical.HITs.SetQuotients as SQ
open import Cubical.HITs.PropositionalTruncation as PT
797 changes: 0 additions & 797 deletions Cubical/Categories/Instances/ZFunctors.agda

This file was deleted.

52 changes: 26 additions & 26 deletions Cubical/Papers/AffineSchemes.agda
Original file line number Diff line number Diff line change
@@ -22,55 +22,55 @@ module Cubical.Papers.AffineSchemes where

-- 2: Background
-- 2.2: Cubical Agda
import Cubical.Foundations.Prelude as Prelude
import Cubical.Foundations.HLevels as HLevels
import Cubical.Foundations.Univalence as Univalence
import Cubical.Data.Sigma as Sigma
import Cubical.HITs.PropositionalTruncation as PT
import Cubical.Algebra.DistLattice.Basis as DistLatticeBasis
import Cubical.HITs.SetQuotients as SQ
import Cubical.Foundations.Prelude as Prelude
import Cubical.Foundations.HLevels as HLevels
import Cubical.Foundations.Univalence as Univalence
import Cubical.Data.Sigma as Sigma
import Cubical.HITs.PropositionalTruncation as PT
import Cubical.Algebra.DistLattice.Basis as DistLatticeBasis
import Cubical.HITs.SetQuotients as SQ


-- 3: Commutative Algebra
-- 3.1: Localization
import Cubical.Algebra.CommRing.Localisation.Base as L
import Cubical.Algebra.CommRing.Localisation.Base as L
module Localization = L.Loc

import Cubical.Algebra.CommRing.Localisation.UniversalProperty as LocalizationUnivProp
import Cubical.Algebra.CommRing.Localisation.InvertingElements as LocalizationInvEl
import Cubical.Algebra.CommAlgebra as R-Algs
import Cubical.Algebra.CommAlgebra.Localisation as LocalizationR-Alg
import Cubical.Algebra.CommRing.Localisation.UniversalProperty as LocalizationUnivProp
import Cubical.Algebra.CommRing.Localisation.InvertingElements as LocalizationInvEl
import Cubical.Algebra.CommAlgebra as R-Algs
import Cubical.Algebra.CommAlgebra.Localisation as LocalizationR-Alg


-- 3.2: The Zariski Lattice
import Cubical.Data.FinData.Base as FiniteTypes
import Cubical.Algebra.Matrix as Matrices
import Cubical.Algebra.CommRing.FGIdeal as FinGenIdeals
import Cubical.Data.FinData.Base as FiniteTypes
import Cubical.Algebra.Matrix as Matrices
import Cubical.Algebra.CommRing.FGIdeal as FinGenIdeals

import Cubical.Algebra.ZariskiLattice.Base as ZLB
import Cubical.AlgebraicGeometry.ZariskiLattice.Base as ZLB
module ZariskiLatDef = ZLB.ZarLat

import Cubical.Algebra.ZariskiLattice.UniversalProperty as ZLUP
import Cubical.AlgebraicGeometry.ZariskiLattice.UniversalProperty as ZLUP
module ZariskiLatUnivProp = ZLUP.ZarLatUniversalProp


-- 4: Category Theory
-- background theory not explicitly mentioned
import Cubical.Categories.Category.Base as CatTheory
import Cubical.Categories.Limits as GeneralLimits
import Cubical.Categories.Limits.RightKan as GeneralKanExtension
import Cubical.Categories.Category.Base as CatTheory
import Cubical.Categories.Limits as GeneralLimits
import Cubical.Categories.Limits.RightKan as GeneralKanExtension

import Cubical.Categories.DistLatticeSheaf.Diagram as SheafDiagShapes
import Cubical.Categories.DistLatticeSheaf.Base as Sheaves
import Cubical.Categories.DistLatticeSheaf.Diagram as SheafDiagShapes
import Cubical.Categories.DistLatticeSheaf.Base as Sheaves

import Cubical.Categories.DistLatticeSheaf.Extension as E
import Cubical.Categories.DistLatticeSheaf.Extension as E
module SheafExtension = E.PreSheafExtension


-- 5: The Structure Sheaf
import Cubical.Categories.Instances.CommAlgebras as R-AlgConstructions
import Cubical.Algebra.CommRing.Localisation.Limit as LocalizationLimit
import Cubical.Algebra.ZariskiLattice.StructureSheaf as StructureSheaf
import Cubical.Categories.Instances.CommAlgebras as R-AlgConstructions
import Cubical.Algebra.CommRing.Localisation.Limit as LocalizationLimit
import Cubical.AlgebraicGeometry.ZariskiLattice.StructureSheaf as StructureSheaf



0 comments on commit 9741a31

Please sign in to comment.