Skip to content

Commit

Permalink
more boilterplate, morphism contructor
Browse files Browse the repository at this point in the history
  • Loading branch information
felixwellen committed Aug 26, 2024
1 parent f368b61 commit 9ce0eff
Show file tree
Hide file tree
Showing 2 changed files with 46 additions and 3 deletions.
15 changes: 12 additions & 3 deletions Cubical/Algebra/CommAlgebra/Base.agda
Original file line number Diff line number Diff line change
Expand Up @@ -138,9 +138,9 @@ module _ {R : CommRing ℓ} where
idCommAlgebraHom = idCAlgHom

compCommAlgebraHom : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} {C : CommAlgebra R ℓ'''}
(g : CommAlgebraHom B C) (f : CommAlgebraHom A B)
(f : CommAlgebraHom A B) (g : CommAlgebraHom B C)
CommAlgebraHom A C
compCommAlgebraHom {A = A} {B = B} {C = C} g f =
compCommAlgebraHom {A = A} {B = B} {C = C} f g =
CommRingHom→CommAlgebraHom (⟨ g ⟩ᵣ→ ∘cr ⟨ f ⟩ᵣ→) pasting
where
opaque
Expand All @@ -160,7 +160,7 @@ module _ {R : CommRing ℓ} where
infixr 9 _∘ca_ -- same as functions
_∘ca_ : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''} {C : CommAlgebra R ℓ'''}
CommAlgebraHom B C CommAlgebraHom A B CommAlgebraHom A C
g ∘ca f = compCommAlgebraHom g f
g ∘ca f = compCommAlgebraHom f g

opaque
CommAlgebraHom≡ :
Expand Down Expand Up @@ -227,6 +227,15 @@ module _ {R : CommRing ℓ} where
CommAlgebra→CommRingStr : (A : CommAlgebra R ℓ') CommRingStr ⟨ A ⟩ₐ
CommAlgebra→CommRingStr A = A .fst .snd

CommAlgebra→Set : (A : CommAlgebra R ℓ') hSet ℓ'
CommAlgebra→Set A = ⟨ A ⟩ₐ , is-set
where open CommRingStr (CommAlgebra→CommRingStr A) using (is-set)

isSetCommAlgebraHom : (A : CommAlgebra R ℓ') (B : CommAlgebra R ℓ'')
isSet (CommAlgebraHom A B)
isSetCommAlgebraHom A B = isSetΣSndProp (isSet→ is-set) (λ _ isPropIsCommAlgebraHom _ _ _)
where open CommRingStr (CommAlgebra→CommRingStr B) using (is-set)

CommAlgebraHom→RingHom : {A : CommAlgebra R ℓ'} {B : CommAlgebra R ℓ''}
CommAlgebraHom A B
RingHom (CommAlgebra→Ring A) (CommAlgebra→Ring B)
Expand Down
34 changes: 34 additions & 0 deletions Cubical/Algebra/CommAlgebra/Properties.agda
Original file line number Diff line number Diff line change
Expand Up @@ -46,3 +46,37 @@ module _ {R : CommRing ℓ} where
f⁻¹ .fst .fst ∘ B .snd .fst ≡⟨ sym ( cong ((f⁻¹ .fst .fst) ∘_) (cong fst (IsCommAlgebraHom.commutes (f .snd)))) ⟩
f⁻¹ .fst .fst ∘ f .fst .fst ∘ A .snd .fst ≡⟨ cong (_∘ A .snd .fst) f⁻¹∘f≡Id ⟩
A .snd .fst ∎

module _
-- Variable generalization would fail below without the module parameters A and B.
{R : CommRing ℓ}
{A : CommAlgebra R ℓ'}
{B : CommAlgebra R ℓ''}
{f : ⟨ A ⟩ₐ ⟨ B ⟩ₐ}
where

open CommAlgebraStr ⦃...⦄
open CommRingStr ⦃...⦄
private instance
_ = CommAlgebra→CommRingStr A
_ = CommAlgebra→CommRingStr B
_ = CommAlgebra→CommAlgebraStr A
_ = CommAlgebra→CommAlgebraStr B

module _
(p1 : f 1r ≡ 1r)
(p+ : (x y : ⟨ A ⟩ₐ) f (x + y) ≡ f x + f y)
(p· : (x y : ⟨ A ⟩ₐ) f (x · y) ≡ f x · f y)
(p⋆ : (r : ⟨ R ⟩) (x : ⟨ A ⟩ₐ) f (r ⋆ x) ≡ r ⋆ f x)
where

makeIsCommAlgebraHom : IsCommAlgebraHom A B f
makeIsCommAlgebraHom .IsCommAlgebraHom.isCommRingHom = makeIsCommRingHom p1 p+ p·
makeIsCommAlgebraHom .IsCommAlgebraHom.commutes =
CommRingHom≡ $ funExt λ r f (A .snd .fst r) ≡⟨ cong f (sym (·IdR _)) ⟩
f ((A .snd .fst r) · 1r) ≡⟨⟩
f (r ⋆ 1r) ≡⟨ p⋆ _ _ ⟩
r ⋆ (f 1r) ≡⟨⟩
(B .snd .fst r) · f 1r ≡⟨ cong ((B .snd .fst r) ·_) p1 ⟩
(B .snd .fst r) · 1r ≡⟨ ·IdR _ ⟩
B .snd .fst r ∎

0 comments on commit 9ce0eff

Please sign in to comment.