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.
strengthen makeCommAlgebraHom
Browse files Browse the repository at this point in the history
MatthiasHu authored and felixwellen committed Jul 24, 2024
1 parent 1352741 commit f9cbbb5
Showing 3 changed files with 17 additions and 9 deletions.
13 changes: 10 additions & 3 deletions Cubical/Algebra/CommAlgebra/Base.agda
Original file line number Diff line number Diff line change
@@ -207,9 +207,9 @@ module _ {R : CommRing ℓ} where
(fPres1 : f 1a ≡ 1a)
(fPres+ : (x y : fst M) f (x + y) ≡ f x + f y)
(fPres· : (x y : fst M) f (x · y) ≡ f x · f y)
(fPres⋆ : (r : fst R) (x : fst M) f (r ⋆ x) ≡ r ⋆ f x)
(fPres⋆1a : (r : fst R) f (r ⋆ 1a) ≡ r ⋆ 1a)
CommAlgebraHom M N
makeCommAlgebraHom f fPres1 fPres+ fPres· fPres⋆ = f , isHom
makeCommAlgebraHom f fPres1 fPres+ fPres· fPres⋆1a = f , isHom
where fPres0 =
f 0a ≡⟨ sym (+IdR _) ⟩
f 0a + 0a ≡⟨ cong (λ u f 0a + u) (sym (+InvR (f 0a))) ⟩
@@ -232,7 +232,14 @@ module _ {R : CommRing ℓ} where
(f ((- x) + x) - f x) ≡⟨ cong (λ u f u - f x) (+InvL x) ⟩
(f 0a - f x) ≡⟨ cong (λ u u - f x) fPres0 ⟩
(0a - f x) ≡⟨ +IdL _ ⟩ (- f x) ∎)
pres⋆ isHom = fPres⋆
pres⋆ isHom r y =
f (r ⋆ y) ≡⟨ cong f (cong (r ⋆_) (sym (·IdL y))) ⟩
f (r ⋆ (1a · y)) ≡⟨ cong f (sym (⋆AssocL r 1a y)) ⟩
f ((r ⋆ 1a) · y) ≡⟨ fPres· (r ⋆ 1a) y ⟩
f (r ⋆ 1a) · f y ≡⟨ cong (_· f y) (fPres⋆1a r) ⟩
(r ⋆ 1a) · f y ≡⟨ ⋆AssocL r 1a (f y) ⟩
r ⋆ (1a · f y) ≡⟨ cong (r ⋆_) (·IdL (f y)) ⟩
r ⋆ f y ∎

isPropIsCommAlgebraHom : (f : fst M fst N) isProp (IsCommAlgebraHom (snd M) f (snd N))
isPropIsCommAlgebraHom f = isPropIsAlgebraHom
4 changes: 2 additions & 2 deletions Cubical/Algebra/CommAlgebra/Instances/Initial.agda
Original file line number Diff line number Diff line change
@@ -64,8 +64,8 @@ module _ (R : CommRing ℓ) where
x * (y * 1a) ≡[ i ]⟨ x * (·IdL (y * 1a) (~ i)) ⟩
x * (1a · (y * 1a)) ≡⟨ sym (⋆AssocL _ _ _) ⟩
(x * 1a) · (y * 1a) ∎)
(λ r x (r · x) * 1a ≡⟨ ⋆Assoc _ _ _
(r * (x * 1a)) ∎)
(λ r (r · _) * 1a ≡⟨ cong (_* 1a) (·IdR r)
(r * 1a) ∎)

initialMapEq : (f : CommAlgebraHom initialCAlg A)
f ≡ initialMap
9 changes: 5 additions & 4 deletions Cubical/Algebra/CommAlgebra/Properties.agda
Original file line number Diff line number Diff line change
@@ -157,10 +157,11 @@ module CommAlgChar (R : CommRing ℓ) {ℓ' : Level} where
instance
_ = snd A
_ = snd B
pres⋆h : r x fst h (fst f r · x) ≡ fst g r · fst h x
pres⋆h r x = fst h (fst f r · x) ≡⟨ pres· (snd h) _ _ ⟩
fst h (fst f r) · fst h x ≡⟨ cong (λ φ fst φ r · fst h x) commDiag ⟩
fst g r · fst h x ∎
pres⋆h : r fst h (fst f r · 1r) ≡ fst g r · 1r
pres⋆h r = fst h (fst f r · 1r) ≡⟨ pres· (snd h) _ _ ⟩
fst h (fst f r) · fst h 1r ≡⟨ cong (λ φ fst φ r · fst h 1r) commDiag ⟩
fst g r · fst h 1r ≡⟨ cong (fst g r ·_) (pres1 (snd h)) ⟩
fst g r · 1r ∎

fromCommAlgebraHom : (A B : CommAlgebra R ℓ') CommAlgebraHom A B
CommRingWithHomHom (fromCommAlg A) (fromCommAlg B)

0 comments on commit f9cbbb5

Please sign in to comment.