Skip to content

Commit d06c432

Browse files
JacquesCarettejamesmckinna
authored andcommitted
Very dependent map (#2373)
* add some 'very dependent' maps to Data.Product. More documentation to come later. * and document * make imports more precise * minimal properties of very-dependen map and zipWith. Structured to make it easy to add more. * whitespace * implement new names and suggestions about using pattern-matching in the type * whitespace in CHANGELOG * small cleanup based on latest round of comments * and fix the names in the comments too. --------- Co-authored-by: jamesmckinna <[email protected]>
1 parent 4749905 commit d06c432

File tree

3 files changed

+95
-5
lines changed

3 files changed

+95
-5
lines changed

CHANGELOG.md

+17
Original file line numberDiff line numberDiff line change
@@ -586,6 +586,23 @@ Additions to existing modules
586586
mapMaybe-↭ : xs ↭ ys → mapMaybe f xs ↭ mapMaybe f ys
587587
```
588588

589+
* Added some very-dependent map and zipWith to `Data.Product`.
590+
```agda
591+
map-Σ : {B : A → Set b} {P : A → Set p} {Q : {x : A} → P x → B x → Set q} →
592+
(f : (x : A) → B x) → (∀ {x} → (y : P x) → Q y (f x)) →
593+
((x , y) : Σ A P) → Σ (B x) (Q y)
594+
595+
map-Σ′ : {B : A → Set b} {P : Set p} {Q : P → Set q} →
596+
(f : (x : A) → B x) → ((x : P) → Q x) → ((x , y) : A × P) → B x × Q y
597+
598+
zipWith : {P : A → Set p} {Q : B → Set q} {R : C → Set r} {S : (x : C) → R x → Set s}
599+
(_∙_ : A → B → C) → (_∘_ : ∀ {x y} → P x → Q y → R (x ∙ y)) →
600+
(_*_ : (x : C) → (y : R x) → S x y) →
601+
((a , p) : Σ A P) → ((b , q) : Σ B Q) →
602+
S (a ∙ b) (p ∘ q)
603+
604+
```
605+
589606
* In `Data.Rational.Properties`:
590607
```agda
591608
1≢0 : 1ℚ ≢ 0ℚ

src/Data/Product.agda

+25-5
Original file line numberDiff line numberDiff line change
@@ -8,20 +8,40 @@
88

99
module Data.Product where
1010

11-
open import Function.Base
12-
open import Level
13-
open import Relation.Nullary.Negation.Core
11+
open import Level using (Level; _⊔_)
12+
open import Relation.Nullary.Negation.Core using (¬_)
1413

1514
private
1615
variable
17-
a b c ℓ : Level
18-
A B : Set a
16+
a b c ℓ p q r s : Level
17+
A B C : Set a
1918

2019
------------------------------------------------------------------------
2120
-- Definition of dependent products
2221

2322
open import Data.Product.Base public
2423

24+
-- These are here as they are not 'basic' but instead "very dependent",
25+
-- i.e. the result type depends on the full input 'point' (x , y).
26+
map-Σ : {B : A Set b} {P : A Set p} {Q : {x : A} P x B x Set q}
27+
(f : (x : A) B x) ( {x} (y : P x) Q y (f x))
28+
((x , y) : Σ A P) Σ (B x) (Q y)
29+
map-Σ f g (x , y) = (f x , g y)
30+
31+
-- This is a "non-dependent" version of map-Σ whereby the input is actually
32+
-- a pair (i.e. _×_ ) but the output type still depends on the input 'point' (x , y).
33+
map-Σ′ : {B : A Set b} {P : Set p} {Q : P Set q}
34+
(f : (x : A) B x) ((x : P) Q x) ((x , y) : A × P) B x × Q y
35+
map-Σ′ f g (x , y) = (f x , g y)
36+
37+
-- This is a generic zipWith for Σ where different functions are applied to each
38+
-- component pair, and recombined.
39+
zipWith : {P : A Set p} {Q : B Set q} {R : C Set r} {S : (x : C) R x Set s}
40+
(_∙_ : A B C) (_∘_ : {x y} P x Q y R (x ∙ y))
41+
(_*_ : (x : C) (y : R x) S x y)
42+
((a , p) : Σ A P) ((b , q) : Σ B Q) S (a ∙ b) (p ∘ q)
43+
zipWith _∙_ _∘_ _*_ (a , p) (b , q) = (a ∙ b) * (p ∘ q)
44+
2545
------------------------------------------------------------------------
2646
-- Negation of existential quantifier
2747

+53
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,53 @@
1+
------------------------------------------------------------------------
2+
-- The Agda standard library
3+
--
4+
-- Properties of 'very dependent' map / zipWith over products
5+
------------------------------------------------------------------------
6+
7+
{-# OPTIONS --cubical-compatible --safe #-}
8+
9+
module Data.Product.Properties.Dependent where
10+
11+
open import Data.Product using (Σ; _×_; _,_; map-Σ; map-Σ′; zipWith)
12+
open import Function.Base using (id; flip)
13+
open import Level using (Level)
14+
open import Relation.Binary.PropositionalEquality.Core
15+
using (_≡_; _≗_; cong₂; refl)
16+
17+
private
18+
variable
19+
a b p q r s : Level
20+
A B C : Set a
21+
22+
------------------------------------------------------------------------
23+
-- map-Σ
24+
25+
module _ {B : A Set b} {P : A Set p} {Q : {x : A} P x B x Set q} where
26+
27+
map-Σ-cong : {f g : (x : A) B x} {h k : {x} (y : P x) Q y (f x)}
28+
( x f x ≡ g x)
29+
( {x} (y : P x) h y ≡ k y)
30+
(v : Σ A P) map-Σ f h v ≡ map-Σ g k v
31+
map-Σ-cong f≗g h≗k (x , y) = cong₂ _,_ (f≗g x) (h≗k y)
32+
33+
------------------------------------------------------------------------
34+
-- map-Σ′
35+
36+
module _ {B : A Set b} {P : Set p} {Q : P Set q} where
37+
38+
map-Σ′-cong : {f g : (x : A) B x} {h k : (x : P) Q x}
39+
( x f x ≡ g x)
40+
((y : P) h y ≡ k y)
41+
(v : A × P) map-Σ′ f h v ≡ map-Σ′ g k v
42+
map-Σ′-cong f≗g h≗k (x , y) = cong₂ _,_ (f≗g x) (h≗k y)
43+
44+
------------------------------------------------------------------------
45+
-- zipWith
46+
47+
module _ {P : A Set p} {Q : B Set q} {R : C Set r} {S : (x : C) R x Set s} where
48+
49+
zipWith-flip : (_∙_ : A B C) (_∘_ : {x y} P x Q y R (x ∙ y))
50+
(_*_ : (x : C) (y : R x) S x y)
51+
{x : Σ A P} {y : Σ B Q}
52+
zipWith _∙_ _∘_ _*_ x y ≡ zipWith (flip _∙_) (flip _∘_) _*_ y x
53+
zipWith-flip _∙_ _∘_ _*_ = refl

0 commit comments

Comments
 (0)