Skip to content

Commit

Permalink
Merge pull request #595 from shlevy/⟨⟩-syntax
Browse files Browse the repository at this point in the history
Add syntax typeclass for types with carriers.
  • Loading branch information
ecavallo authored Oct 22, 2021
2 parents 367f3af + a675929 commit 9f57beb
Show file tree
Hide file tree
Showing 3 changed files with 23 additions and 3 deletions.
11 changes: 8 additions & 3 deletions Cubical/Foundations/Structure.agda
Original file line number Diff line number Diff line change
Expand Up @@ -3,6 +3,7 @@ module Cubical.Foundations.Structure where

open import Cubical.Core.Everything
open import Cubical.Foundations.Prelude
open import Cubical.Syntax.⟨⟩

private
variable
Expand All @@ -21,9 +22,13 @@ typ = fst
str : (A : TypeWithStr ℓ S) S (typ A)
str = snd

-- Alternative notation for typ
⟨_⟩ : TypeWithStr ℓ S Type ℓ
⟨_⟩ = typ
instance
TypeWithStr-has-⟨⟩ : {ℓ S} has-⟨⟩ (TypeWithStr {ℓ' = ℓ'} ℓ S)
⟨_⟩ ⦃ TypeWithStr-has-⟨⟩ ⦄ = typ

-- Allow users to avoid importing the syntax module directly for
-- backwards compatibility.
open import Cubical.Syntax.⟨⟩ using (⟨_⟩) public

-- An S-structure should have a notion of S-homomorphism, or rather S-isomorphism.
-- This will be implemented by a function ι : StrEquiv S ℓ'
Expand Down
3 changes: 3 additions & 0 deletions Cubical/README.agda
Original file line number Diff line number Diff line change
Expand Up @@ -68,3 +68,6 @@ import Cubical.Reflection.Everything

-- Displayed univalent graphs
import Cubical.Displayed.Everything

-- Syntax typeclasses
import Cubical.Syntax.Everything
12 changes: 12 additions & 0 deletions Cubical/Syntax/⟨⟩.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,12 @@
{-# OPTIONS --safe #-}
module Cubical.Syntax.⟨⟩ where

open import Cubical.Core.Primitives

-- A syntax typeclass for types which contain a "carrier" type in the
-- sence of an algebraic structure.
record has-⟨⟩ {ℓᵢ ℓc} (Instance : Type ℓᵢ) : Type (ℓ-max ℓᵢ (ℓ-suc ℓc)) where
field
⟨_⟩ : Instance Type ℓc

open has-⟨⟩ ⦃ … ⦄ public

0 comments on commit 9f57beb

Please sign in to comment.