Skip to content

Commit dca8857

Browse files
WhatisRTjamesmckinnagallaisJacquesCarette
authored
Add some type relations and isomorphisms (#2419)
* Add some type relations/isomorphisms * Remove some explicitly bound variables in type signatures * Add new functions to CHANGELOG * Update src/Data/Product/Function/Dependent/Propositional.agda Name change * Update CHANGELOG.md Name change inherited from source update * Update src/Function/Related/TypeIsomorphisms.agda Remove redundant parentheses * Update TypeIsomorphisms.agda @JacquesCarette @jamesmckinna @gallais switched the prime discipline on the distributive laws for products and sums * Update CHANGELOG.md * [ fix ] whitespace violation * use the more polymorphic one here --------- Co-authored-by: jamesmckinna <[email protected]> Co-authored-by: Guillaume Allais <[email protected]> Co-authored-by: Jacques Carette <[email protected]>
1 parent a18e51b commit dca8857

File tree

3 files changed

+60
-13
lines changed

3 files changed

+60
-13
lines changed

CHANGELOG.md

Lines changed: 16 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -19,6 +19,8 @@ Bug-fixes
1919
Non-backwards compatible changes
2020
--------------------------------
2121

22+
In `Function.Related.TypeIsomorphisms`, the unprimed versions are more level polymorphic; and the primed versions retain `Level` homogeneous types for the `Semiring` axioms to hold.
23+
2224
Minor improvements
2325
------------------
2426

@@ -353,6 +355,11 @@ Additions to existing modules
353355
m≤pred[n]⇒suc[m]≤n : .{{NonZero n}} → m ≤ pred n → suc m ≤ n
354356
```
355357

358+
* In `Data.Product.Function.Dependent.Propositional`:
359+
```agda
360+
congˡ : ∀ {k} → (∀ {x} → A x ∼[ k ] B x) → Σ I A ∼[ k ] Σ I B
361+
```
362+
356363
* New lemmas in `Data.Rational.Properties`:
357364
```agda
358365
nonNeg+nonNeg⇒nonNeg : ∀ p .{{_ : NonNegative p}} q .{{_ : NonNegative q}} → NonNegative (p + q)
@@ -389,6 +396,15 @@ Additions to existing modules
389396
_≡?_ : DecidableEquality (Vec A n)
390397
```
391398

399+
* In `Function.Related.TypeIsomorphisms`:
400+
```agda
401+
Σ-distribˡ-⊎ : (∃ λ a → P a ⊎ Q a) ↔ (∃ P ⊎ ∃ Q)
402+
Σ-distribʳ-⊎ : (Σ (A ⊎ B) P) ↔ (Σ A (P ∘ inj₁) ⊎ Σ B (P ∘ inj₂))
403+
×-distribˡ-⊎ : (A × (B ⊎ C)) ↔ (A × B ⊎ A × C)
404+
×-distribʳ-⊎ : ((A ⊎ B) × C) ↔ (A × C ⊎ B × C)
405+
∃-≡ : ∀ (P : A → Set b) {x} → P x ↔ (∃[ y ] y ≡ x × P y)
406+
```
407+
392408
* In `Relation.Binary.Bundles`:
393409
```agda
394410
record DecPreorder c ℓ₁ ℓ₂ : Set (suc (c ⊔ ℓ₁ ⊔ ℓ₂))

src/Data/Product/Function/Dependent/Propositional.agda

Lines changed: 4 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -17,7 +17,7 @@ open import Function.Related.Propositional
1717
implication; reverseImplication; equivalence; injection;
1818
reverseInjection; leftInverse; surjection; bijection)
1919
open import Function.Base using (_$_; _∘_; _∘′_)
20-
open import Function.Properties.Inverse using (↔⇒↠; ↔⇒⟶; ↔⇒⟵; ↔-sym; ↔⇒↩)
20+
open import Function.Properties.Inverse using (↔⇒↠; ↔⇒⟶; ↔⇒⟵; ↔-sym; ↔⇒↩; refl)
2121
open import Function.Properties.RightInverse using (↩⇒↪; ↪⇒↩)
2222
open import Function.Properties.Inverse.HalfAdjointEquivalence
2323
using (↔⇒≃; _≃_; ≃⇒↔)
@@ -316,3 +316,6 @@ cong {B = B} {k = reverseInjection} I↔J A↢B = Σ-↣ (↔-sym I↔J) (swap
316316
cong {B = B} {k = leftInverse} I↔J A↩B = ↩⇒↪ (Σ-↩ (↔⇒↩ (↔-sym I↔J)) (↪⇒↩ (swap-coercions {k = leftInverse} B I↔J A↩B)))
317317
cong {k = surjection} I↔J A↠B = Σ-↠ (↔⇒↠ I↔J) A↠B
318318
cong {k = bijection} I↔J A↔B = Σ-↔ I↔J A↔B
319+
320+
congˡ : {k} ( {x} A x ∼[ k ] B x) Σ I A ∼[ k ] Σ I B
321+
congˡ = cong (refl _)

src/Function/Related/TypeIsomorphisms.agda

Lines changed: 40 additions & 12 deletions
Original file line numberDiff line numberDiff line change
@@ -15,7 +15,7 @@ open import Axiom.Extensionality.Propositional using (Extensionality)
1515
open import Data.Bool.Base using (true; false)
1616
open import Data.Empty.Polymorphic using (⊥; ⊥-elim)
1717
open import Data.Product.Base as Product
18-
using (_×_; Σ; curry; uncurry; _,_; -,_; <_,_>; proj₁; proj₂; ∃₂; ∃)
18+
using (_×_; Σ; curry; uncurry; _,_; -,_; <_,_>; proj₁; proj₂; ∃₂; ∃; ∃-syntax)
1919
open import Data.Product.Function.NonDependent.Propositional
2020
open import Data.Sum.Base as Sum
2121
open import Data.Sum.Properties using (swap-involutive)
@@ -106,26 +106,46 @@ private
106106
⊎-identity ℓ = ⊎-identityˡ ℓ , ⊎-identityʳ ℓ
107107

108108
------------------------------------------------------------------------
109-
-- Properties of × and ⊎
109+
-- Properties of Σ and ⊎
110110

111-
-- × distributes over ⊎
111+
-- Σ distributes over ⊎
112112

113-
×-distribˡ-⊎ : _DistributesOverˡ_ {ℓ = ℓ} _↔_ _×_ _⊎_
114-
×-distribˡ-⊎ ℓ _ _ _ = mk↔ₛ′
113+
Σ-distribˡ-⊎ : {P : A Set a} {Q : A Set b}
114+
(∃ λ a P a ⊎ Q a) ↔ (∃ P ⊎ ∃ Q)
115+
Σ-distribˡ-⊎ = mk↔ₛ′
115116
(uncurry λ x [ inj₁ ∘′ (x ,_) , inj₂ ∘′ (x ,_) ]′)
116117
[ Product.map₂ inj₁ , Product.map₂ inj₂ ]′
117118
[ (λ _ refl) , (λ _ refl) ]
118119
(uncurry λ _ [ (λ _ refl) , (λ _ refl) ])
119120

120-
×-distribʳ-⊎ : _DistributesOverʳ_ {ℓ = ℓ} _↔_ _×_ _⊎_
121-
×-distribʳ-⊎ ℓ _ _ _ = mk↔ₛ′
122-
(uncurry [ curry inj₁ , curry inj₂ ]′)
123-
[ Product.map₁ inj₁ , Product.map₁ inj₂ ]′
121+
Σ-distribʳ-⊎ : {P : A ⊎ B Set c}
122+
(Σ (A ⊎ B) P) ↔ (Σ A (P ∘ inj₁) ⊎ Σ B (P ∘ inj₂))
123+
Σ-distribʳ-⊎ = mk↔ₛ′
124+
(uncurry [ curry inj₁ , curry inj₂ ])
125+
[ Product.dmap inj₁ id , Product.dmap inj₂ id ]
124126
[ (λ _ refl) , (λ _ refl) ]
125127
(uncurry [ (λ _ _ refl) , (λ _ _ refl) ])
126128

127-
×-distrib-⊎ : _DistributesOver_ {ℓ = ℓ} _↔_ _×_ _⊎_
128-
×-distrib-⊎ ℓ = ×-distribˡ-⊎ ℓ , ×-distribʳ-⊎ ℓ
129+
------------------------------------------------------------------------
130+
-- Properties of × and ⊎
131+
132+
-- × distributes over ⊎
133+
-- primed variants are less level polymorphic
134+
135+
×-distribˡ-⊎ : (A × (B ⊎ C)) ↔ (A × B ⊎ A × C)
136+
×-distribˡ-⊎ = Σ-distribˡ-⊎
137+
138+
×-distribˡ-⊎′ : _DistributesOverˡ_ {ℓ = ℓ} _↔_ _×_ _⊎_
139+
×-distribˡ-⊎′ ℓ _ _ _ = ×-distribˡ-⊎
140+
141+
×-distribʳ-⊎ : ((A ⊎ B) × C) ↔ (A × C ⊎ B × C)
142+
×-distribʳ-⊎ = Σ-distribʳ-⊎
143+
144+
×-distribʳ-⊎′ : _DistributesOverʳ_ {ℓ = ℓ} _↔_ _×_ _⊎_
145+
×-distribʳ-⊎′ ℓ _ _ _ = ×-distribʳ-⊎
146+
147+
×-distrib-⊎′ : _DistributesOver_ {ℓ = ℓ} _↔_ _×_ _⊎_
148+
×-distrib-⊎′ ℓ = ×-distribˡ-⊎′ ℓ , ×-distribʳ-⊎′ ℓ
129149

130150
------------------------------------------------------------------------
131151
-- ⊥, ⊤, _×_ and _⊎_ form a commutative semiring
@@ -228,7 +248,7 @@ private
228248
×-⊎-isCommutativeSemiring k ℓ = isCommutativeSemiringˡ record
229249
{ +-isCommutativeMonoid = ⊎-isCommutativeMonoid k ℓ
230250
; *-isCommutativeMonoid = ×-isCommutativeMonoid k ℓ
231-
; distribʳ = λ A B C ↔⇒ (×-distribʳ-⊎ ℓ A B C)
251+
; distribʳ = λ _ _ _ ↔⇒ ×-distribʳ-⊎
232252
; zeroˡ = ↔⇒ ∘ ×-zeroˡ ℓ
233253
}
234254

@@ -332,3 +352,11 @@ True↔ ( true because [p]) irr =
332352
mk↔ₛ′ (λ _ invert [p]) (λ _ _) (irr _) (λ _ refl)
333353
True↔ (false because ofⁿ ¬p) _ =
334354
mk↔ₛ′ (λ()) (invert (ofⁿ ¬p)) (λ x flip contradiction ¬p x) (λ ())
355+
356+
------------------------------------------------------------------------
357+
-- Relating a predicate to an existentially quantified one with the
358+
-- restriction that the quantified variable is equal to the given one
359+
360+
∃-≡ : (P : A Set b) {x} P x ↔ (∃[ y ] y ≡ x × P y)
361+
∃-≡ P {x} = mk↔ₛ′ (λ Px x , refl , Px) (λ where (_ , refl , Py) Py)
362+
(λ where (_ , refl , _) refl) (λ where _ refl)

0 commit comments

Comments
 (0)