Skip to content

Commit

Permalink
Merge pull request #976 from agda/issue6368-compat
Browse files Browse the repository at this point in the history
[ re agda/#6368 ] Add blanket instance for TypeWithStr
  • Loading branch information
ecavallo authored Jan 27, 2023
2 parents 545e873 + f755e38 commit c011356
Show file tree
Hide file tree
Showing 3 changed files with 4 additions and 5 deletions.
2 changes: 0 additions & 2 deletions Cubical/Algebra/Ring/Properties.agda
Original file line number Diff line number Diff line change
Expand Up @@ -252,8 +252,6 @@ module RingHomTheory {R : Ring ℓ} {S : Ring ℓ'} (φ : RingHom R S) where
open IsRingHom (φ .snd)
private
instance
_ = R
_ = S
_ = snd R
_ = snd S
f = fst φ
Expand Down
2 changes: 0 additions & 2 deletions Cubical/Algebra/Ring/Quotient.agda
Original file line number Diff line number Diff line change
Expand Up @@ -215,15 +215,13 @@ module UniversalProperty (R : Ring ℓ) (I : IdealsIn R) where
Iₛ = fst I
private
instance
_ = R
_ = snd R

module _ {S : Ring ℓ'} (φ : RingHom R S) where
open IsRingHom
open RingHomTheory φ
private
instance
_ = S
_ = snd S
f = fst φ
module φ = IsRingHom (snd φ)
Expand Down
5 changes: 4 additions & 1 deletion Cubical/Foundations/Structure.agda
Original file line number Diff line number Diff line change
Expand Up @@ -25,6 +25,10 @@ str = snd
⟨_⟩ : TypeWithStr ℓ S Type ℓ
⟨_⟩ = typ

instance
mkTypeWithStr : {ℓ} {S : Type ℓ Type ℓ'} {X} {{S X}} TypeWithStr ℓ S
mkTypeWithStr {{i}} = _ , i

-- An S-structure should have a notion of S-homomorphism, or rather S-isomorphism.
-- This will be implemented by a function ι : StrEquiv S ℓ'
-- that gives us for any two types with S-structure (X , s) and (Y , t) a family:
Expand All @@ -41,4 +45,3 @@ EquivAction {ℓ} S = {X Y : Type ℓ} → X ≃ Y → S X ≃ S Y
EquivAction→StrEquiv : {S : Type ℓ Type ℓ''}
EquivAction S StrEquiv S ℓ''
EquivAction→StrEquiv α (X , s) (Y , t) e = equivFun (α e) s ≡ t

0 comments on commit c011356

Please sign in to comment.