Skip to content

Commit

Permalink
Structure presheaf (#699)
Browse files Browse the repository at this point in the history
* use improved ringsolver

* delete one more line

* code from old branch

* port to CommRings

* our lemma at the right place

* give up for today

* give up for today

* find a new strategy

* give up

* simplify

* prepare for presheaf

* Add boiler plate code needed

* done

* clean up
  • Loading branch information
mzeuner authored Jan 25, 2022
1 parent e1ba6d8 commit 6371279
Show file tree
Hide file tree
Showing 24 changed files with 843 additions and 87 deletions.
1 change: 1 addition & 0 deletions Cubical/Algebra/Algebra.agda
Original file line number Diff line number Diff line change
Expand Up @@ -2,3 +2,4 @@
module Cubical.Algebra.Algebra where

open import Cubical.Algebra.Algebra.Base public
open import Cubical.Algebra.Algebra.Properties public
58 changes: 11 additions & 47 deletions Cubical/Algebra/Algebra/Base.agda
Original file line number Diff line number Diff line change
Expand Up @@ -156,27 +156,18 @@ open IsAlgebraHom
AlgebraHom : {R : Ring ℓ} (M : Algebra R ℓ') (N : Algebra R ℓ'') Type (ℓ-max ℓ (ℓ-max ℓ' ℓ''))
AlgebraHom M N = Σ[ f ∈ (⟨ M ⟩ ⟨ N ⟩) ] IsAlgebraHom (M .snd) f (N .snd)

idAlgHom : {R : Ring ℓ} {A : Algebra R ℓ'} AlgebraHom A A
fst idAlgHom x = x
pres0 (snd idAlgHom) = refl
pres1 (snd idAlgHom) = refl
pres+ (snd idAlgHom) x y = refl
pres· (snd idAlgHom) x y = refl
pres- (snd idAlgHom) x = refl
pres⋆ (snd idAlgHom) r x = refl

IsAlgebraEquiv : {R : Ring ℓ} {A B : Type ℓ'}
IsAlgebraEquiv : {R : Ring ℓ} {A : Type ℓ'} {B : Type ℓ''}
(M : AlgebraStr R A) (e : A ≃ B) (N : AlgebraStr R B)
Type (ℓ-max ℓ ℓ')
Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'')
IsAlgebraEquiv M e N = IsAlgebraHom M (e .fst) N

AlgebraEquiv : {R : Ring ℓ} (M N : Algebra R ℓ') Type (ℓ-max ℓ ℓ')
AlgebraEquiv : {R : Ring ℓ} (M : Algebra R ℓ') (N : Algebra R ℓ'') Type (ℓ-max (ℓ-max ℓ ℓ') ℓ'')
AlgebraEquiv M N = Σ[ e ∈ ⟨ M ⟩ ≃ ⟨ N ⟩ ] IsAlgebraEquiv (M .snd) e (N .snd)

_$a_ : {R : Ring ℓ} {A : Algebra R ℓ'} {B : Algebra R ℓ''} AlgebraHom A B ⟨ A ⟩ ⟨ B ⟩
f $a x = fst f x

AlgebraEquiv→AlgebraHom : {R : Ring ℓ} {A B : Algebra R ℓ'}
AlgebraEquiv→AlgebraHom : {R : Ring ℓ} {A : Algebra R ℓ'} {B : Algebra R ℓ''}
AlgebraEquiv A B AlgebraHom A B
AlgebraEquiv→AlgebraHom (e , eIsHom) = e .fst , eIsHom

Expand Down Expand Up @@ -218,6 +209,9 @@ isSetAlgebraEquiv : {R : Ring ℓ} (M N : Algebra R ℓ')
isSetAlgebraEquiv M N = isSetΣ (isOfHLevel≃ 2 (isSetAlgebra M) (isSetAlgebra N))
λ _ isProp→isSet (isPropIsAlgebraHom _ _ _ _)

AlgebraHom≡ : {R : Ring ℓ} {A B : Algebra R ℓ'} {φ ψ : AlgebraHom A B} fst φ ≡ fst ψ φ ≡ ψ
AlgebraHom≡ = Σ≡Prop λ f isPropIsAlgebraHom _ _ f _

𝒮ᴰ-Algebra : (R : Ring ℓ) DUARel (𝒮-Univ ℓ') (AlgebraStr R) (ℓ-max ℓ ℓ')
𝒮ᴰ-Algebra R =
𝒮ᴰ-Record (𝒮-Univ _) (IsAlgebraEquiv {R = R})
Expand All @@ -239,41 +233,11 @@ isSetAlgebraEquiv M N = isSetΣ (isOfHLevel≃ 2 (isSetAlgebra M) (isSetAlgebra
AlgebraPath : {R : Ring ℓ} (A B : Algebra R ℓ') (AlgebraEquiv A B) ≃ (A ≡ B)
AlgebraPath {R = R} = ∫ (𝒮ᴰ-Algebra R) .UARel.ua

compIsAlgebraHom : {R : Ring ℓ} {A : Algebra R ℓ'} {B : Algebra R ℓ''} {C : Algebra R ℓ'''}
{g : ⟨ B ⟩ ⟨ C ⟩} {f : ⟨ A ⟩ ⟨ B ⟩}
IsAlgebraHom (B .snd) g (C .snd)
IsAlgebraHom (A .snd) f (B .snd)
IsAlgebraHom (A .snd) (g ∘ f) (C .snd)
compIsAlgebraHom {g = g} {f} gh fh .pres0 = cong g (fh .pres0) ∙ gh .pres0
compIsAlgebraHom {g = g} {f} gh fh .pres1 = cong g (fh .pres1) ∙ gh .pres1
compIsAlgebraHom {g = g} {f} gh fh .pres+ x y = cong g (fh .pres+ x y) ∙ gh .pres+ (f x) (f y)
compIsAlgebraHom {g = g} {f} gh fh .pres· x y = cong g (fh .pres· x y) ∙ gh .pres· (f x) (f y)
compIsAlgebraHom {g = g} {f} gh fh .pres- x = cong g (fh .pres- x) ∙ gh .pres- (f x)
compIsAlgebraHom {g = g} {f} gh fh .pres⋆ r x = cong g (fh .pres⋆ r x) ∙ gh .pres⋆ r (f x)

_∘a_ : {R : Ring ℓ} {A : Algebra R ℓ'} {B : Algebra R ℓ''} {C : Algebra R ℓ'''}
AlgebraHom B C AlgebraHom A B AlgebraHom A C
_∘a_ g f .fst = g .fst ∘ f .fst
_∘a_ g f .snd = compIsAlgebraHom (g .snd) (f .snd)

module AlgebraTheory (R : Ring ℓ) (A : Algebra R ℓ') where
open RingStr (snd R) renaming (_+_ to _+r_ ; _·_ to _·r_)
open AlgebraStr (A .snd)

0-actsNullifying : (x : ⟨ A ⟩) 0r ⋆ x ≡ 0a
0-actsNullifying x =
let idempotent-+ = 0r ⋆ x ≡⟨ cong (λ u u ⋆ x) (sym (RingTheory.0Idempotent R)) ⟩
(0r +r 0r) ⋆ x ≡⟨ ⋆-ldist 0r 0r x ⟩
(0r ⋆ x) + (0r ⋆ x) ∎
in RingTheory.+Idempotency→0 (Algebra→Ring A) (0r ⋆ x) idempotent-+

⋆Dist· : (x y : ⟨ R ⟩) (a b : ⟨ A ⟩) (x ·r y) ⋆ (a · b) ≡ (x ⋆ a) · (y ⋆ b)
⋆Dist· x y a b = (x ·r y) ⋆ (a · b) ≡⟨ ⋆-rassoc _ _ _ ⟩
a · ((x ·r y) ⋆ b) ≡⟨ cong (a ·_) (⋆-assoc _ _ _) ⟩
a · (x ⋆ (y ⋆ b)) ≡⟨ sym (⋆-rassoc _ _ _) ⟩
x ⋆ (a · (y ⋆ b)) ≡⟨ sym (⋆-lassoc _ _ _) ⟩
(x ⋆ a) · (y ⋆ b) ∎
uaAlgebra : {R : Ring ℓ} {A B : Algebra R ℓ'} AlgebraEquiv A B A ≡ B
uaAlgebra {A = A} {B = B} = equivFun (AlgebraPath A B)

isGroupoidAlgebra : {R : Ring ℓ} isGroupoid (Algebra R ℓ')
isGroupoidAlgebra _ _ = isOfHLevelRespectEquiv 2 (AlgebraPath _ _) (isSetAlgebraEquiv _ _)

-- Smart constructor for ring homomorphisms
-- that infers the other equations from pres1, pres+, and pres·
Expand Down
170 changes: 170 additions & 0 deletions Cubical/Algebra/Algebra/Properties.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,170 @@
{-# OPTIONS --safe #-}
module Cubical.Algebra.Algebra.Properties where

open import Cubical.Foundations.Prelude
open import Cubical.Foundations.Function
open import Cubical.Foundations.Equiv
open import Cubical.Foundations.Equiv.HalfAdjoint
open import Cubical.Foundations.HLevels
open import Cubical.Foundations.Isomorphism
open import Cubical.Foundations.Univalence
open import Cubical.Foundations.Transport
open import Cubical.Foundations.SIP
open import Cubical.Foundations.GroupoidLaws
open import Cubical.Foundations.Path

open import Cubical.Data.Sigma

open import Cubical.Structures.Axioms
open import Cubical.Structures.Auto
open import Cubical.Structures.Macro
open import Cubical.Algebra.Module
open import Cubical.Algebra.Ring
open import Cubical.Algebra.AbGroup
open import Cubical.Algebra.Group
open import Cubical.Algebra.Monoid
open import Cubical.Algebra.Algebra.Base

open Iso

private
variable
ℓ ℓ' ℓ'' ℓ''' : Level




module AlgebraTheory (R : Ring ℓ) (A : Algebra R ℓ') where
open RingStr (snd R) renaming (_+_ to _+r_ ; _·_ to _·r_)
open AlgebraStr (A .snd)

0-actsNullifying : (x : ⟨ A ⟩) 0r ⋆ x ≡ 0a
0-actsNullifying x =
let idempotent-+ = 0r ⋆ x ≡⟨ cong (λ u u ⋆ x) (sym (RingTheory.0Idempotent R)) ⟩
(0r +r 0r) ⋆ x ≡⟨ ⋆-ldist 0r 0r x ⟩
(0r ⋆ x) + (0r ⋆ x) ∎
in RingTheory.+Idempotency→0 (Algebra→Ring A) (0r ⋆ x) idempotent-+

⋆Dist· : (x y : ⟨ R ⟩) (a b : ⟨ A ⟩) (x ·r y) ⋆ (a · b) ≡ (x ⋆ a) · (y ⋆ b)
⋆Dist· x y a b = (x ·r y) ⋆ (a · b) ≡⟨ ⋆-rassoc _ _ _ ⟩
a · ((x ·r y) ⋆ b) ≡⟨ cong (a ·_) (⋆-assoc _ _ _) ⟩
a · (x ⋆ (y ⋆ b)) ≡⟨ sym (⋆-rassoc _ _ _) ⟩
x ⋆ (a · (y ⋆ b)) ≡⟨ sym (⋆-lassoc _ _ _) ⟩
(x ⋆ a) · (y ⋆ b) ∎


module AlgebraHoms {R : Ring ℓ} where
open IsAlgebraHom

idAlgebraHom : (A : Algebra R ℓ') AlgebraHom A A
fst (idAlgebraHom A) x = x
pres0 (snd (idAlgebraHom A)) = refl
pres1 (snd (idAlgebraHom A)) = refl
pres+ (snd (idAlgebraHom A)) x y = refl
pres· (snd (idAlgebraHom A)) x y = refl
pres- (snd (idAlgebraHom A)) x = refl
pres⋆ (snd (idAlgebraHom A)) r x = refl

compIsAlgebraHom : {A : Algebra R ℓ'} {B : Algebra R ℓ''} {C : Algebra R ℓ'''}
{g : ⟨ B ⟩ ⟨ C ⟩} {f : ⟨ A ⟩ ⟨ B ⟩}
IsAlgebraHom (B .snd) g (C .snd)
IsAlgebraHom (A .snd) f (B .snd)
IsAlgebraHom (A .snd) (g ∘ f) (C .snd)
compIsAlgebraHom {g = g} {f} gh fh .pres0 = cong g (fh .pres0) ∙ gh .pres0
compIsAlgebraHom {g = g} {f} gh fh .pres1 = cong g (fh .pres1) ∙ gh .pres1
compIsAlgebraHom {g = g} {f} gh fh .pres+ x y = cong g (fh .pres+ x y) ∙ gh .pres+ (f x) (f y)
compIsAlgebraHom {g = g} {f} gh fh .pres· x y = cong g (fh .pres· x y) ∙ gh .pres· (f x) (f y)
compIsAlgebraHom {g = g} {f} gh fh .pres- x = cong g (fh .pres- x) ∙ gh .pres- (f x)
compIsAlgebraHom {g = g} {f} gh fh .pres⋆ r x = cong g (fh .pres⋆ r x) ∙ gh .pres⋆ r (f x)

compAlgebraHom : {A : Algebra R ℓ'} {B : Algebra R ℓ''} {C : Algebra R ℓ'''}
AlgebraHom A B AlgebraHom B C AlgebraHom A C
compAlgebraHom f g .fst = g .fst ∘ f .fst
compAlgebraHom f g .snd = compIsAlgebraHom (g .snd) (f .snd)

syntax compAlgebraHom f g = g ∘a f

compIdAlgebraHom : {A B : Algebra R ℓ'} (φ : AlgebraHom A B) compAlgebraHom (idAlgebraHom A) φ ≡ φ
compIdAlgebraHom φ = AlgebraHom≡ refl

idCompAlgebraHom : {A B : Algebra R ℓ'} (φ : AlgebraHom A B) compAlgebraHom φ (idAlgebraHom B) ≡ φ
idCompAlgebraHom φ = AlgebraHom≡ refl

compAssocAlgebraHom : {A B C D : Algebra R ℓ'}
: AlgebraHom A B) (ψ : AlgebraHom B C) (χ : AlgebraHom C D)
compAlgebraHom (compAlgebraHom φ ψ) χ ≡ compAlgebraHom φ (compAlgebraHom ψ χ)
compAssocAlgebraHom _ _ _ = AlgebraHom≡ refl


module AlgebraEquivs {R : Ring ℓ} where
open IsAlgebraHom
open AlgebraHoms

compIsAlgebraEquiv : {A : Algebra R ℓ'} {B : Algebra R ℓ''} {C : Algebra R ℓ'''}
{g : ⟨ B ⟩ ≃ ⟨ C ⟩} {f : ⟨ A ⟩ ≃ ⟨ B ⟩}
IsAlgebraEquiv (B .snd) g (C .snd)
IsAlgebraEquiv (A .snd) f (B .snd)
IsAlgebraEquiv (A .snd) (compEquiv f g) (C .snd)
compIsAlgebraEquiv {g = g} {f} gh fh = compIsAlgebraHom {g = g .fst} {f .fst} gh fh

compAlgebraEquiv : {A : Algebra R ℓ'} {B : Algebra R ℓ''} {C : Algebra R ℓ'''}
AlgebraEquiv A B AlgebraEquiv B C AlgebraEquiv A C
fst (compAlgebraEquiv f g) = compEquiv (f .fst) (g .fst)
snd (compAlgebraEquiv f g) = compIsAlgebraEquiv {g = g .fst} {f = f .fst} (g .snd) (f .snd)



-- the Algebra version of uaCompEquiv
module AlgebraUAFunctoriality {R : Ring ℓ} where
open AlgebraStr
open AlgebraEquivs

Algebra≡ : (A B : Algebra R ℓ') (
Σ[ p ∈ ⟨ A ⟩ ≡ ⟨ B ⟩ ]
Σ[ q0 ∈ PathP (λ i p i) (0a (snd A)) (0a (snd B)) ]
Σ[ q1 ∈ PathP (λ i p i) (1a (snd A)) (1a (snd B)) ]
Σ[ r+ ∈ PathP (λ i p i p i p i) (_+_ (snd A)) (_+_ (snd B)) ]
Σ[ r· ∈ PathP (λ i p i p i p i) (_·_ (snd A)) (_·_ (snd B)) ]
Σ[ s- ∈ PathP (λ i p i p i) (-_ (snd A)) (-_ (snd B)) ]
Σ[ s⋆ ∈ PathP (λ i ⟨ R ⟩ p i p i) (_⋆_ (snd A)) (_⋆_ (snd B)) ]
PathP (λ i IsAlgebra R (q0 i) (q1 i) (r+ i) (r· i) (s- i) (s⋆ i)) (isAlgebra (snd A))
(isAlgebra (snd B)))
≃ (A ≡ B)
Algebra≡ A B = isoToEquiv theIso
where
open Iso
theIso : Iso _ _
fun theIso (p , q0 , q1 , r+ , r· , s- , s⋆ , t) i = p i
, algebrastr (q0 i) (q1 i) (r+ i) (r· i) (s- i) (s⋆ i) (t i)
inv theIso x = cong ⟨_⟩ x , cong (0a ∘ snd) x , cong (1a ∘ snd) x
, cong (_+_ ∘ snd) x , cong (_·_ ∘ snd) x , cong (-_ ∘ snd) x , cong (_⋆_ ∘ snd) x
, cong (isAlgebra ∘ snd) x
rightInv theIso _ = refl
leftInv theIso _ = refl

caracAlgebra≡ : {A B : Algebra R ℓ'} (p q : A ≡ B) cong ⟨_⟩ p ≡ cong ⟨_⟩ q p ≡ q
caracAlgebra≡ {A = A} {B = B} p q P =
sym (transportTransport⁻ (ua (Algebra≡ A B)) p)
∙∙ cong (transport (ua (Algebra≡ A B))) helper
∙∙ transportTransport⁻ (ua (Algebra≡ A B)) q
where
helper : transport (sym (ua (Algebra≡ A B))) p ≡ transport (sym (ua (Algebra≡ A B))) q
helper = Σ≡Prop
(λ _ isPropΣ
(isOfHLevelPathP' 1 (is-set (snd B)) _ _)
λ _ isPropΣ (isOfHLevelPathP' 1 (is-set (snd B)) _ _)
λ _ isPropΣ (isOfHLevelPathP' 1 (isSetΠ2 λ _ _ is-set (snd B)) _ _)
λ _ isPropΣ (isOfHLevelPathP' 1 (isSetΠ2 λ _ _ is-set (snd B)) _ _)
λ _ isPropΣ (isOfHLevelPathP' 1 (isSetΠ λ _ is-set (snd B)) _ _)
λ _ isPropΣ (isOfHLevelPathP' 1 (isSetΠ2 λ _ _ is-set (snd B)) _ _)
λ _ isOfHLevelPathP 1 (isPropIsAlgebra _ _ _ _ _ _ _) _ _)
(transportRefl (cong ⟨_⟩ p) ∙ P ∙ sym (transportRefl (cong ⟨_⟩ q)))

uaCompAlgebraEquiv : {A B C : Algebra R ℓ'} (f : AlgebraEquiv A B) (g : AlgebraEquiv B C)
uaAlgebra (compAlgebraEquiv f g) ≡ uaAlgebra f ∙ uaAlgebra g
uaCompAlgebraEquiv f g = caracAlgebra≡ _ _ (
cong ⟨_⟩ (uaAlgebra (compAlgebraEquiv f g))
≡⟨ uaCompEquiv _ _ ⟩
cong ⟨_⟩ (uaAlgebra f) ∙ cong ⟨_⟩ (uaAlgebra g)
≡⟨ sym (cong-∙ ⟨_⟩ (uaAlgebra f) (uaAlgebra g)) ⟩
cong ⟨_⟩ (uaAlgebra f ∙ uaAlgebra g) ∎)
1 change: 1 addition & 0 deletions Cubical/Algebra/CommAlgebra.agda
Original file line number Diff line number Diff line change
Expand Up @@ -2,3 +2,4 @@
module Cubical.Algebra.CommAlgebra where

open import Cubical.Algebra.CommAlgebra.Base public
open import Cubical.Algebra.CommAlgebra.Properties public
21 changes: 14 additions & 7 deletions Cubical/Algebra/CommAlgebra/Base.agda
Original file line number Diff line number Diff line change
Expand Up @@ -23,7 +23,7 @@ open import Cubical.Reflection.RecordEquiv

private
variable
ℓ ℓ' : Level
ℓ ℓ' ℓ'' ℓ''' : Level

record IsCommAlgebra (R : CommRing ℓ) {A : Type ℓ'}
(0a : A) (1a : A)
Expand Down Expand Up @@ -142,24 +142,28 @@ module _ {R : CommRing ℓ} where
·Assoc⋆ ⋆DistR ⋆DistL ⋆Lid ⋆Assoc·)


IsCommAlgebraEquiv : {A B : Type ℓ'}
IsCommAlgebraEquiv : {A : Type ℓ'} {B : Type ℓ''}
(M : CommAlgebraStr R A) (e : A ≃ B) (N : CommAlgebraStr R B)
Type (ℓ-max ℓ ℓ')
Type _
IsCommAlgebraEquiv M e N =
IsAlgebraHom (CommAlgebraStr→AlgebraStr M) (e .fst) (CommAlgebraStr→AlgebraStr N)

CommAlgebraEquiv : (M N : CommAlgebra R ℓ') Type (ℓ-max ℓ ℓ')
CommAlgebraEquiv : (M : CommAlgebra R ℓ') (N : CommAlgebra R ℓ'') Type _
CommAlgebraEquiv M N = Σ[ e ∈ ⟨ M ⟩ ≃ ⟨ N ⟩ ] IsCommAlgebraEquiv (M .snd) e (N .snd)

IsCommAlgebraHom : {A B : Type ℓ'}
IsCommAlgebraHom : {A : Type ℓ'} {B : Type ℓ''}
(M : CommAlgebraStr R A) (f : A B) (N : CommAlgebraStr R B)
Type (ℓ-max ℓ ℓ')
Type _
IsCommAlgebraHom M f N =
IsAlgebraHom (CommAlgebraStr→AlgebraStr M) f (CommAlgebraStr→AlgebraStr N)

CommAlgebraHom : (M N : CommAlgebra R ℓ') Type (ℓ-max ℓ ℓ')
CommAlgebraHom : (M : CommAlgebra R ℓ') (N : CommAlgebra R ℓ'') Type _
CommAlgebraHom M N = Σ[ f ∈ (⟨ M ⟩ ⟨ N ⟩) ] IsCommAlgebraHom (M .snd) f (N .snd)

CommAlgebraEquiv→CommAlgebraHom : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''}
CommAlgebraEquiv A B CommAlgebraHom A B
CommAlgebraEquiv→CommAlgebraHom (e , eIsHom) = e .fst , eIsHom

module _ {M N : CommAlgebra R ℓ'} where
open CommAlgebraStr {{...}}
open IsAlgebraHom
Expand Down Expand Up @@ -239,5 +243,8 @@ isPropIsCommAlgebra R _ _ _ _ _ _ =
CommAlgebraPath : (R : CommRing ℓ) (A B : CommAlgebra R ℓ') (CommAlgebraEquiv A B) ≃ (A ≡ B)
CommAlgebraPath R = ∫ (𝒮ᴰ-CommAlgebra R) .UARel.ua

uaCommAlgebra : {R : CommRing ℓ} {A B : CommAlgebra R ℓ'} CommAlgebraEquiv A B A ≡ B
uaCommAlgebra {R = R} {A = A} {B = B} = equivFun (CommAlgebraPath R A B)

isGroupoidCommAlgebra : {R : CommRing ℓ} isGroupoid (CommAlgebra R ℓ')
isGroupoidCommAlgebra A B = isOfHLevelRespectEquiv 2 (CommAlgebraPath _ _ _) (isSetAlgebraEquiv _ _)
1 change: 1 addition & 0 deletions Cubical/Algebra/CommAlgebra/Instances/FreeCommAlgebra.agda
Original file line number Diff line number Diff line change
Expand Up @@ -427,6 +427,7 @@ inducedHom : {R : CommRing ℓ} {I : Type ℓ'} (A : CommAlgebra R ℓ'')
inducedHom A φ = Theory.inducedHom A φ

module _ {R : CommRing ℓ} {A B : CommAlgebra R ℓ''} where
open AlgebraHoms
A′ = CommAlgebra→Algebra A
B′ = CommAlgebra→Algebra B
R′ = (CommRing→Ring R)
Expand Down
14 changes: 11 additions & 3 deletions Cubical/Algebra/CommAlgebra/Localisation.agda
Original file line number Diff line number Diff line change
Expand Up @@ -21,6 +21,7 @@ open import Cubical.Algebra.CommRing.Base
open import Cubical.Algebra.CommRing.Properties
open import Cubical.Algebra.CommRing.Localisation.Base
open import Cubical.Algebra.CommRing.Localisation.UniversalProperty
open import Cubical.Algebra.CommRing.Localisation.InvertingElements
open import Cubical.Algebra.Ring
open import Cubical.Algebra.Algebra
open import Cubical.Algebra.CommAlgebra.Base
Expand Down Expand Up @@ -125,7 +126,12 @@ module AlgLoc (R' : CommRing ℓ)
isContrHomS⁻¹RS⁻¹R = S⁻¹RHasAlgUniversalProp S⁻¹RAsCommAlg S⋆1⊆S⁻¹Rˣ



-- the special case of localizing at a single element
R[1/_]AsCommAlgebra : {R : CommRing ℓ} ⟨ R ⟩ CommAlgebra R ℓ
R[1/_]AsCommAlgebra {R = R} f = S⁻¹RAsCommAlg [ f ⁿ|n≥0] (powersFormMultClosedSubset f)
where
open AlgLoc R
open InvertingElementsBase R

module AlgLocTwoSubsets (R' : CommRing ℓ)
(S₁ : ℙ (fst R')) (S₁MultClosedSubset : isMultClosedSubset R' S₁)
Expand All @@ -147,6 +153,8 @@ module AlgLocTwoSubsets (R' : CommRing ℓ)
; isContrHomS⁻¹RS⁻¹R to isContrHomS₂⁻¹RS₂⁻¹R)

open IsAlgebraHom
open AlgebraHoms
open CommAlgebraHoms
open CommAlgebraStr ⦃...⦄

private
Expand All @@ -171,10 +179,10 @@ module AlgLocTwoSubsets (R' : CommRing ℓ)
χ₂ : CommAlgebraHom S₂⁻¹RAsCommAlg S₁⁻¹RAsCommAlg
χ₂ = S₂⁻¹RHasAlgUniversalProp S₁⁻¹RAsCommAlg S₂⊆S₁⁻¹Rˣ .fst

χ₁∘χ₂≡id : χ₁ ∘a χ₂ ≡ idAlgHom
χ₁∘χ₂≡id : χ₁ ∘a χ₂ ≡ idCommAlgebraHom _
χ₁∘χ₂≡id = isContr→isProp isContrHomS₂⁻¹RS₂⁻¹R _ _

χ₂∘χ₁≡id : χ₂ ∘a χ₁ ≡ idAlgHom
χ₂∘χ₁≡id : χ₂ ∘a χ₁ ≡ idCommAlgebraHom _
χ₂∘χ₁≡id = isContr→isProp isContrHomS₁⁻¹RS₁⁻¹R _ _

IsoS₁⁻¹RS₂⁻¹R : Iso S₁⁻¹R S₂⁻¹R
Expand Down
Loading

0 comments on commit 6371279

Please sign in to comment.