Skip to content

Commit

Permalink
Fix Allable/Anyable + instances for List⁺
Browse files Browse the repository at this point in the history
  • Loading branch information
omelkonian committed Nov 26, 2024
1 parent 70a6c7a commit 2d7cb5c
Show file tree
Hide file tree
Showing 5 changed files with 13 additions and 7 deletions.
6 changes: 3 additions & 3 deletions Class/Allable/Core.agda
Original file line number Diff line number Diff line change
Expand Up @@ -2,12 +2,12 @@ module Class.Allable.Core where

open import Class.Prelude

record Allable (F : Set Set ℓ) : Set (lsuc ℓ) where
field All : {A} (A Set) F A Set
record Allable (F : Type Type ℓ) : Typeω where
field All : {A} (A Type ℓ′) F A Type (ℓ ⊔ ℓ′)

∀∈-syntax = All
∀∈-syntax′ = All
¬∀∈-syntax = λ {A} P ¬_ ∘ All {A} P
¬∀∈-syntax = λ {A} {ℓ′} (P : A Type ℓ′) ¬_ ∘ All P
¬∀∈-syntax′ = ¬∀∈-syntax
infix 2 ∀∈-syntax ∀∈-syntax′ ¬∀∈-syntax ¬∀∈-syntax′
syntax ∀∈-syntax P xs = ∀[∈ xs ] P
Expand Down
3 changes: 3 additions & 0 deletions Class/Allable/Instance.agda
Original file line number Diff line number Diff line change
Expand Up @@ -17,6 +17,9 @@ instance
Allable-Maybe : Allable {ℓ} Maybe
Allable-Maybe .All = M.All

Allable-List⁺ : Allable {ℓ} List⁺
Allable-List⁺ .All P = All P ∘ toList

private
open import Class.Decidable
open import Class.HasOrder
Expand Down
6 changes: 3 additions & 3 deletions Class/Anyable/Core.agda
Original file line number Diff line number Diff line change
Expand Up @@ -2,12 +2,12 @@ module Class.Anyable.Core where

open import Class.Prelude

record Anyable (F : Set Set ℓ) : Set (lsuc ℓ) where
field Any : {A} (A Set) F A Set
record Anyable (F : Type Type ℓ) : Typeω where
field Any : {A} (A Type ℓ′) F A Type (ℓ ⊔ ℓ′)

∃∈-syntax = Any
∃∈-syntax′ = Any
∄∈-syntax = λ {A} P ¬_ ∘ Any {A} P
∄∈-syntax = λ {A} {ℓ′} (P : A Type ℓ′) ¬_ ∘ Any P
∄∈-syntax′ = ∄∈-syntax
infix 2 ∃∈-syntax ∃∈-syntax′ ∄∈-syntax ∄∈-syntax′
syntax ∃∈-syntax P xs = ∃[∈ xs ] P
Expand Down
3 changes: 3 additions & 0 deletions Class/Anyable/Instance.agda
Original file line number Diff line number Diff line change
Expand Up @@ -17,6 +17,9 @@ instance
Anyable-Maybe : Anyable {ℓ} Maybe
Anyable-Maybe .Any = M.Any

Anyable-List⁺ : Anyable {ℓ} List⁺
Anyable-List⁺ .Any P = Any P ∘ toList

private
open import Class.Decidable
open import Class.HasOrder
Expand Down
2 changes: 1 addition & 1 deletion Class/Prelude.agda
Original file line number Diff line number Diff line change
Expand Up @@ -37,7 +37,7 @@ open import Data.Maybe public
open import Data.List public
using (List; []; _∷_; [_]; map; _++_; foldr; concat; concatMap)
open import Data.List.NonEmpty public
using (List⁺; _∷_; _⁺++⁺_; foldr₁)
using (List⁺; _∷_; _⁺++⁺_; foldr₁; toList)
open import Data.Vec public
using (Vec; []; _∷_)
open import Data.These public
Expand Down

0 comments on commit 2d7cb5c

Please sign in to comment.