Skip to content
New issue

Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.

By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.

Already on GitHub? Sign in to your account

CommAlgebras as CommRingHoms #1145

Draft
wants to merge 83 commits into
base: master
Choose a base branch
from
Draft
Show file tree
Hide file tree
Changes from all commits
Commits
Show all changes
83 commits
Select commit Hold shift + click to select a range
d464874
move univalence code
felixwellen Jul 29, 2024
9054103
start...
felixwellen Jul 29, 2024
6bd6176
boilerplate
felixwellen Jul 30, 2024
acde70a
fix auto-univalence
felixwellen May 14, 2024
2d9b18f
fix FreeCommAlg
felixwellen May 29, 2024
16ba911
fix subalgebra
felixwellen May 30, 2024
b1bac5c
fix FPAlgebra
felixwellen May 30, 2024
487fe8d
copy Evans approach for Groups
felixwellen Jun 18, 2024
a33924c
refactor
felixwellen Jun 19, 2024
ef58e50
refactor
felixwellen Jun 19, 2024
90bbfac
no-eta for algebra structures
felixwellen Jun 19, 2024
83e9919
opaque quotients, fix unicity
felixwellen Jul 26, 2024
bbda88b
fast tc for quotient comm rings
felixwellen Jul 29, 2024
35274b5
uniform indentation
felixwellen Jul 29, 2024
41bc2a3
cleanup /remove flag
felixwellen Jul 29, 2024
a498039
improve names
felixwellen Jul 29, 2024
3e57fdd
clean up
felixwellen Jul 29, 2024
79f932c
WIP: deactivate problematic code
felixwellen Jul 29, 2024
34fc5c9
refactor/simplify
felixwellen Jun 20, 2024
76f8c22
WIP
felixwellen Jul 26, 2024
8a0c069
start with stuff on the new comm alg
felixwellen Jul 26, 2024
5801fd3
Revert "Auxiliary commit to revert individual files from c5e74aa95d26…
felixwellen Jul 29, 2024
2725661
Quotients
felixwellen Jul 29, 2024
9b9e0ac
update comm ring hom interface
felixwellen Jul 30, 2024
b83bf0f
more boilerplate
felixwellen Jul 30, 2024
a07d7d7
kernels
felixwellen Jul 30, 2024
bd52821
syntax for new algebras
felixwellen Jul 30, 2024
3a9a310
more on comm ring quotients
felixwellen Jul 30, 2024
4a9842f
fix
felixwellen Jul 30, 2024
f1433a8
algebra quotients
felixwellen Jul 30, 2024
d8b07b4
more opaque definitions, more use of set-level equalities
felixwellen Jul 31, 2024
55d58c6
commRingHom fixes
felixwellen Aug 1, 2024
4bacbaa
unit comm algebra
felixwellen Jul 31, 2024
194e456
quotients
felixwellen Jul 31, 2024
62d9af2
no-eta for the equations of a ring
felixwellen Aug 1, 2024
b4475a4
WIP: Univalence
felixwellen Aug 1, 2024
24582d2
equality characterization with help from Evan
felixwellen Aug 1, 2024
58d2e89
WIP
felixwellen Aug 1, 2024
78a6beb
typevariate polynomials as comm ring
felixwellen Aug 2, 2024
bb224c5
polynomials as comm algebra
felixwellen Aug 2, 2024
6667098
remove opaque structure again
felixwellen Aug 2, 2024
f13c4c1
Revert "WIP: deactivate problematic code"
felixwellen Aug 2, 2024
6052d37
make one example more meaningful
felixwellen Aug 2, 2024
3079c8d
univalence for comm algebras with Evans help
felixwellen Aug 2, 2024
527861f
various fixes, continue with typevariate polynomials
felixwellen Aug 2, 2024
631e9bc
induce a com ring hom
felixwellen Aug 2, 2024
2f43b22
Comm ring of polynomials on a general variable type and its universal…
felixwellen Aug 2, 2024
57c7b16
improve interface
felixwellen Aug 3, 2024
b576f32
more interface, more opaque
felixwellen Aug 3, 2024
a19fbd5
fix
felixwellen Aug 3, 2024
5089c7f
start with localization, fixes, more opaque things
felixwellen Aug 3, 2024
e9f0616
make ring quotients' isRing opaque instead of no-eta
felixwellen Aug 3, 2024
2de9d4f
update to new interface for comm ring homs
felixwellen Aug 4, 2024
94770d3
update to new interface for comm ring homs
felixwellen Aug 4, 2024
74d51ab
stub for polynomials, fix
felixwellen Aug 4, 2024
3573a4c
finish localization
felixwellen Aug 4, 2024
414b2f2
start with OnCoproduct, change default notation for R[I]
felixwellen Aug 4, 2024
0b76ab7
universal property of polynomials as rings
felixwellen Aug 7, 2024
1c71ebd
move old CommAlgebra
felixwellen Aug 7, 2024
33c060b
category structure
felixwellen Aug 7, 2024
5ec0ba4
fix ref
felixwellen Aug 8, 2024
acfd747
fix/move algebra as module
felixwellen Aug 8, 2024
0f5053e
refs
felixwellen Aug 8, 2024
367e8bd
remove old CommAlgebra.agda, move all refs
felixwellen Aug 8, 2024
71d44b9
fix
felixwellen Aug 8, 2024
7cc50a6
make the new CommAlgebra the default
felixwellen Aug 8, 2024
f21b3a1
remove no-eta for structure
felixwellen Aug 8, 2024
703e717
start with fp algebras
felixwellen Aug 8, 2024
5384862
small additions
felixwellen Aug 8, 2024
6e12f6c
WIP: start with FP algebra examples
felixwellen Aug 8, 2024
1f44a50
change def, fix thinghs
felixwellen Aug 8, 2024
a95079a
fix localization, adapt univalence, rearrange base
felixwellen Aug 9, 2024
4ba4d01
fix quotients
felixwellen Aug 9, 2024
b2f7072
fix
felixwellen Aug 9, 2024
89ecaa2
fix
felixwellen Aug 23, 2024
7293a0d
solver examples with new comm algebras
felixwellen Aug 26, 2024
bbd32fd
fix
felixwellen Aug 26, 2024
9286b26
more boilterplate, morphism contructor
felixwellen Aug 26, 2024
14c6d4c
polynomial ring is fp
felixwellen Aug 30, 2024
ed48931
add more calcualtions
felixwellen Aug 30, 2024
c6f62d4
port initial algebra
felixwellen Sep 30, 2024
bc7a252
fix comment
felixwellen Sep 30, 2024
7917b85
WIP
felixwellen Oct 10, 2024
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Loading
Jump to
Jump to file
Failed to load files.
Loading
Diff view
Diff view
54 changes: 12 additions & 42 deletions Cubical/Algebra/Algebra/Base.agda
Original file line number Diff line number Diff line change
Expand Up @@ -10,11 +10,6 @@ open import Cubical.Foundations.SIP

open import Cubical.Data.Sigma

open import Cubical.Displayed.Base
open import Cubical.Displayed.Auto
open import Cubical.Displayed.Record
open import Cubical.Displayed.Universe

open import Cubical.Reflection.RecordEquiv

open import Cubical.Algebra.Monoid
Expand All @@ -36,7 +31,7 @@ record IsAlgebra (R : Ring ℓ) {A : Type ℓ'}
(0a 1a : A) (_+_ _·_ : A → A → A) (-_ : A → A)
(_⋆_ : ⟨ R ⟩ → A → A) : Type (ℓ-max ℓ ℓ') where

constructor isalgebra
no-eta-equality

open RingStr (snd R) using (1r) renaming (_+_ to _+r_; _·_ to _·r_)

Expand All @@ -59,8 +54,6 @@ unquoteDecl IsAlgebraIsoΣ = declareRecordIsoΣ IsAlgebraIsoΣ (quote IsAlgebra)

record AlgebraStr (R : Ring ℓ) (A : Type ℓ') : Type (ℓ-max ℓ ℓ') where

constructor algebrastr

field
0a : A
1a : A
Expand All @@ -72,6 +65,8 @@ record AlgebraStr (R : Ring ℓ) (A : Type ℓ') : Type (ℓ-max ℓ ℓ') where

open IsAlgebra isAlgebra public

unquoteDecl AlgebraStrIsoΣ = declareRecordIsoΣ AlgebraStrIsoΣ (quote AlgebraStr)

Algebra : (R : Ring ℓ) → ∀ ℓ' → Type (ℓ-max ℓ (ℓ-suc ℓ'))
Algebra R ℓ' = Σ[ A ∈ Type ℓ' ] AlgebraStr R A

Expand Down Expand Up @@ -208,13 +203,15 @@ isPropIsAlgebra R _ _ _ _ _ _ = let open IsLeftModule in
isPropIsAlgebraHom : (R : Ring ℓ) {A : Type ℓ'} {B : Type ℓ''}
(AS : AlgebraStr R A) (f : A → B) (BS : AlgebraStr R B)
→ isProp (IsAlgebraHom AS f BS)
isPropIsAlgebraHom R AS f BS = isOfHLevelRetractFromIso 1 IsAlgebraHomIsoΣ
(isProp×5 (is-set _ _)
(is-set _ _)
(isPropΠ2 λ _ _ → is-set _ _)
(isPropΠ2 λ _ _ → is-set _ _)
(isPropΠ λ _ → is-set _ _)
(isPropΠ2 λ _ _ → is-set _ _))
isPropIsAlgebraHom R AS f BS =
isOfHLevelRetractFromIso 1
IsAlgebraHomIsoΣ
(isProp×5 (is-set _ _)
(is-set _ _)
(isPropΠ2 λ _ _ → is-set _ _)
(isPropΠ2 λ _ _ → is-set _ _)
(isPropΠ λ _ → is-set _ _)
(isPropΠ2 λ _ _ → is-set _ _))
where
open AlgebraStr BS

Expand All @@ -237,33 +234,6 @@ isSetAlgebraEquiv M N = isSetΣ (isOfHLevel≃ 2 M.is-set N.is-set)
AlgebraHom≡ : {φ ψ : 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})
(fields:
data[ 0a ∣ nul ∣ pres0 ]
data[ 1a ∣ nul ∣ pres1 ]
data[ _+_ ∣ bin ∣ pres+ ]
data[ _·_ ∣ bin ∣ pres· ]
data[ -_ ∣ autoDUARel _ _ ∣ pres- ]
data[ _⋆_ ∣ autoDUARel _ _ ∣ pres⋆ ]
prop[ isAlgebra ∣ (λ _ _ → isPropIsAlgebra _ _ _ _ _ _ _) ])
where
open AlgebraStr

-- faster with some sharing
nul = autoDUARel (𝒮-Univ _) (λ A → A)
bin = autoDUARel (𝒮-Univ _) (λ A → A → A → A)

AlgebraPath : (A B : Algebra R ℓ') → (AlgebraEquiv A B) ≃ (A ≡ B)
AlgebraPath {R = R} = ∫ (𝒮ᴰ-Algebra R) .UARel.ua

uaAlgebra : AlgebraEquiv A B → A ≡ B
uaAlgebra {A = A} {B = B} = equivFun (AlgebraPath A B)

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

-- Smart constructor for algebra homomorphisms
-- that infers the other equations from pres1, pres+, pres·, and pres⋆

Expand Down
69 changes: 13 additions & 56 deletions Cubical/Algebra/Algebra/Properties.agda
Original file line number Diff line number Diff line change
Expand Up @@ -10,8 +10,7 @@ open import Cubical.Foundations.HLevels
open import Cubical.Foundations.GroupoidLaws
open import Cubical.Foundations.Path
open import Cubical.Foundations.Transport
open import Cubical.Foundations.Univalence
open import Cubical.Foundations.SIP
open import Cubical.Foundations.SIP using (⟨_⟩)

open import Cubical.Data.Sigma

Expand Down Expand Up @@ -185,57 +184,15 @@ module AlgebraEquivs where
(λ g → isPropIsAlgebraHom _ (B .snd) g (C .snd))
(isoOnTypes .leftInv (g .fst))

-- the Algebra version of uaCompEquiv
module AlgebraUAFunctoriality 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≡ : (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 : (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) ∎)
isSetAlgebraStr : (A : Type ℓ') → isSet (AlgebraStr R A)
isSetAlgebraStr A =
let open AlgebraStr
in isOfHLevelSucIfInhabited→isOfHLevelSuc 1 λ str →
isOfHLevelRetractFromIso 2 AlgebraStrIsoΣ $
isSetΣ (str .is-set) λ _ →
isSetΣ (str .is-set) λ _ →
isSetΣ (isSet→ (isSet→ (str .is-set))) λ _ →
isSetΣ (isSet→ (isSet→ (str .is-set))) λ _ →
isSetΣ (isSet→ (str .is-set)) (λ _ →
isSetΣSndProp (isSet→ (isSet→ (str .is-set))) λ _ →
isPropIsAlgebra _ _ _ _ _ _ _)
86 changes: 86 additions & 0 deletions Cubical/Algebra/Algebra/Univalence.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,86 @@
{-# OPTIONS --safe #-}
module Cubical.Algebra.Algebra.Univalence where

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

open import Cubical.Functions.Embedding

open import Cubical.Data.Sigma

open import Cubical.Displayed.Base
open import Cubical.Displayed.Auto
open import Cubical.Displayed.Record
open import Cubical.Displayed.Universe

open import Cubical.Algebra.Ring.Base
open import Cubical.Algebra.Algebra.Base
open import Cubical.Algebra.Algebra.Properties

private
variable
ℓ ℓ' ℓ'' ℓ''' : Level
R : Ring ℓ
A B C D : Algebra R ℓ

open IsAlgebraHom

𝒮ᴰ-Algebra : (R : Ring ℓ) → DUARel (𝒮-Univ ℓ') (AlgebraStr R) (ℓ-max ℓ ℓ')
𝒮ᴰ-Algebra R =
𝒮ᴰ-Record (𝒮-Univ _) (IsAlgebraEquiv {R = R})
(fields:
data[ 0a ∣ nul ∣ pres0 ]
data[ 1a ∣ nul ∣ pres1 ]
data[ _+_ ∣ bin ∣ pres+ ]
data[ _·_ ∣ bin ∣ pres· ]
data[ -_ ∣ autoDUARel _ _ ∣ pres- ]
data[ _⋆_ ∣ autoDUARel _ _ ∣ pres⋆ ]
prop[ isAlgebra ∣ (λ A ΣA →
isPropIsAlgebra
R
(snd (fst (fst (fst (fst (fst ΣA))))))
(snd (fst (fst (fst (fst ΣA)))))
(snd (fst (fst (fst ΣA))))
(snd (fst (fst ΣA)))
(snd (fst ΣA))
(snd ΣA)) ])
where
open AlgebraStr

-- faster with some sharing
nul = autoDUARel (𝒮-Univ _) (λ A → A)
bin = autoDUARel (𝒮-Univ _) (λ A → A → A → A)

AlgebraPath : (A B : Algebra R ℓ') → (AlgebraEquiv A B) ≃ (A ≡ B)
AlgebraPath {R = R} = ∫ (𝒮ᴰ-Algebra R) .UARel.ua

uaAlgebra : AlgebraEquiv A B → A ≡ B
uaAlgebra {A = A} {B = B} = equivFun (AlgebraPath A B)

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

-- the Algebra version of uaCompEquiv
module AlgebraUAFunctoriality where
open AlgebraStr
open AlgebraEquivs

caracAlgebra≡ : (p q : A ≡ B) → cong ⟨_⟩ p ≡ cong ⟨_⟩ q → p ≡ q
caracAlgebra≡ _ _ α =
isEmbedding→Inj (iso→isEmbedding (invIso ΣPathIsoPathΣ)) _ _ $
Σ≡Prop (λ _ → isOfHLevelPathP' 1 (isSetAlgebraStr _) _ _) α

uaCompAlgebraEquiv : (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) ∎)
5 changes: 5 additions & 0 deletions Cubical/Algebra/CommAlgebra/AsModule.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,5 @@
{-# OPTIONS --safe #-}
module Cubical.Algebra.CommAlgebra.AsModule where

open import Cubical.Algebra.CommAlgebra.AsModule.Base public
open import Cubical.Algebra.CommAlgebra.AsModule.Properties public
Loading
Loading