Skip to content

Commit 2e4061d

Browse files
guilhermehasguialvaresMatthewDaggitt
authored
Added pointwise and catmaybe in Lists (#2222)
* added pointwise and catmaybe * added pointwise to any * added cat maybe all * changed pointwise definition * changed files name * fixed changelogs and properties * changed Any solution * changed pointwise to catMaybe --------- Co-authored-by: Guilherme <[email protected]> Co-authored-by: MatthewDaggitt <[email protected]>
1 parent cc7d32d commit 2e4061d

File tree

4 files changed

+44
-3
lines changed

4 files changed

+44
-3
lines changed

CHANGELOG.md

+12
Original file line numberDiff line numberDiff line change
@@ -27,7 +27,19 @@ New modules
2727
Additions to existing modules
2828
-----------------------------
2929

30+
* In `Data.Maybe.Relation.Binary.Pointwise`:
31+
```agda
32+
pointwise⊆any : Pointwise R (just x) ⊆ Any (R x)
33+
```
34+
35+
* In `Data.List.Relation.Unary.All.Properties`:
36+
```agda
37+
All-catMaybes⁺ : All (Maybe.All P) xs → All P (catMaybes xs)
38+
Any-catMaybes⁺ : All (Maybe.Any P) xs → All P (catMaybes xs)
39+
```
40+
3041
* In `Data.List.Relation.Unary.AllPairs.Properties`:
3142
```
43+
catMaybes⁺ : AllPairs (Pointwise R) xs → AllPairs R (catMaybes xs)
3244
tabulate⁺-< : (i < j → R (f i) (f j)) → AllPairs R (tabulate f)
3345
```

src/Data/List/Relation/Unary/All/Properties.agda

+13-1
Original file line numberDiff line numberDiff line change
@@ -28,7 +28,8 @@ import Data.List.Relation.Binary.Equality.Setoid as ListEq using (_≋_; []; _
2828
open import Data.List.Relation.Binary.Pointwise.Base using (Pointwise; []; _∷_)
2929
open import Data.List.Relation.Binary.Subset.Propositional using (_⊆_)
3030
open import Data.Maybe.Base as Maybe using (Maybe; just; nothing)
31-
open import Data.Maybe.Relation.Unary.All as Maybe using (just; nothing)
31+
open import Data.Maybe.Relation.Unary.All as Maybe using (just; nothing; fromAny)
32+
open import Data.Maybe.Relation.Unary.Any as Maybe using (just)
3233
open import Data.Nat.Base using (zero; suc; s≤s; _<_; z<s; s<s)
3334
open import Data.Nat.Properties using (≤-refl; m≤n⇒m≤1+n)
3435
open import Data.Product.Base as Prod using (_×_; _,_; uncurry; uncurry′)
@@ -388,6 +389,17 @@ mapMaybe⁺ {xs = x ∷ xs} {f = f} (px ∷ pxs) with f x
388389
... | just v with px
389390
... | just pv = pv ∷ mapMaybe⁺ pxs
390391

392+
------------------------------------------------------------------------
393+
-- catMaybes
394+
395+
All-catMaybes⁺ : All (Maybe.All P) xs All P (catMaybes xs)
396+
All-catMaybes⁺ [] = []
397+
All-catMaybes⁺ (just px ∷ pxs) = px ∷ All-catMaybes⁺ pxs
398+
All-catMaybes⁺ (nothing ∷ pxs) = All-catMaybes⁺ pxs
399+
400+
Any-catMaybes⁺ : All (Maybe.Any P) xs All P (catMaybes xs)
401+
Any-catMaybes⁺ = All-catMaybes⁺ ∘ All.map fromAny
402+
391403
------------------------------------------------------------------------
392404
-- _++_
393405

src/Data/List/Relation/Unary/AllPairs/Properties.agda

+14-2
Original file line numberDiff line numberDiff line change
@@ -10,9 +10,11 @@ module Data.List.Relation.Unary.AllPairs.Properties where
1010

1111
open import Data.List.Base hiding (any)
1212
open import Data.List.Relation.Unary.All as All using (All; []; _∷_)
13-
import Data.List.Relation.Unary.All.Properties as All
13+
open import Data.List.Relation.Unary.All.Properties as All using (Any-catMaybes⁺)
1414
open import Data.List.Relation.Unary.AllPairs as AllPairs using (AllPairs; []; _∷_)
1515
open import Data.Bool.Base using (true; false)
16+
open import Data.Maybe using (Maybe; nothing; just)
17+
open import Data.Maybe.Relation.Binary.Pointwise using (pointwise⊆any; Pointwise)
1618
open import Data.Fin.Base as F using (Fin)
1719
open import Data.Fin.Properties using (suc-injective; <⇒≢)
1820
open import Data.Nat.Base using (zero; suc; _<_; z≤n; s≤s; z<s; s<s)
@@ -22,7 +24,7 @@ open import Level using (Level)
2224
open import Relation.Binary.Core using (Rel)
2325
open import Relation.Binary.Bundles using (DecSetoid)
2426
open import Relation.Binary.PropositionalEquality.Core using (_≢_)
25-
open import Relation.Unary using (Pred; Decidable)
27+
open import Relation.Unary using (Pred; Decidable; _⊆_)
2628
open import Relation.Nullary.Decidable using (does)
2729

2830
private
@@ -136,3 +138,13 @@ module _ {R : Rel A ℓ} {P : Pred A p} (P? : Decidable P) where
136138
filter⁺ {x ∷ xs} (x∉xs ∷ xs!) with does (P? x)
137139
... | false = filter⁺ xs!
138140
... | true = All.filter⁺ P? x∉xs ∷ filter⁺ xs!
141+
142+
------------------------------------------------------------------------
143+
-- catMaybes
144+
145+
module _ {R : Rel A ℓ} where
146+
147+
catMaybes⁺ : {xs : List (Maybe A)} AllPairs (Pointwise R) xs AllPairs R (catMaybes xs)
148+
catMaybes⁺ {xs = []} [] = []
149+
catMaybes⁺ {xs = nothing ∷ _} (x∼xs ∷ pxs) = catMaybes⁺ pxs
150+
catMaybes⁺ {xs = just x ∷ xs} (x∼xs ∷ pxs) = Any-catMaybes⁺ (All.map pointwise⊆any x∼xs) ∷ catMaybes⁺ pxs

src/Data/Maybe/Relation/Binary/Pointwise.agda

+5
Original file line numberDiff line numberDiff line change
@@ -11,13 +11,15 @@ module Data.Maybe.Relation.Binary.Pointwise where
1111
open import Level
1212
open import Data.Product.Base using (∃; _×_; -,_; _,_)
1313
open import Data.Maybe.Base using (Maybe; just; nothing)
14+
open import Data.Maybe.Relation.Unary.Any using (Any; just)
1415
open import Function.Bundles using (_⇔_; mk⇔)
1516
open import Relation.Binary.Core using (REL; Rel; _⇒_)
1617
open import Relation.Binary.Bundles using (Setoid; DecSetoid)
1718
open import Relation.Binary.Definitions using (Reflexive; Sym; Trans; Decidable)
1819
open import Relation.Binary.Structures using (IsEquivalence; IsDecEquivalence)
1920
open import Relation.Binary.PropositionalEquality.Core as P using (_≡_)
2021
open import Relation.Nullary
22+
open import Relation.Unary using (_⊆_)
2123
import Relation.Nullary.Decidable as Dec
2224

2325
------------------------------------------------------------------------
@@ -93,6 +95,9 @@ module _ {a r} {A : Set a} {R : Rel A r} where
9395
; _≟_ = dec R._≟_
9496
} where module R = IsDecEquivalence R-isDecEquivalence
9597

98+
pointwise⊆any : {x} Pointwise R (just x) ⊆ Any (R x)
99+
pointwise⊆any (just Rxy) = just Rxy
100+
96101
module _ {c ℓ} where
97102

98103
setoid : Setoid c ℓ Setoid c (c ⊔ ℓ)

0 commit comments

Comments
 (0)