Skip to content

Commit 82438db

Browse files
Saransh-cppMatthewDaggitt
authored andcommitted
Migrate _↔̇_ to the new Function hierarchy + more arrows for indexed sets (#1982)
1 parent 4b008d6 commit 82438db

File tree

2 files changed

+56
-0
lines changed

2 files changed

+56
-0
lines changed

CHANGELOG.md

Lines changed: 6 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1624,6 +1624,12 @@ New modules
16241624
```
16251625
Algebra.Properties.KleeneAlgebra
16261626
```
1627+
1628+
* Relations on indexed sets
1629+
```
1630+
Function.Indexed.Bundles
1631+
```
1632+
16271633
Other minor changes
16281634
-------------------
16291635

src/Function/Indexed/Bundles.agda

Lines changed: 50 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,50 @@
1+
------------------------------------------------------------------------
2+
-- The Agda standard library
3+
--
4+
-- Operations on Relations for Indexed sets
5+
------------------------------------------------------------------------
6+
7+
{-# OPTIONS --cubical-compatible --safe #-}
8+
9+
module Function.Indexed.Bundles where
10+
11+
open import Relation.Unary using (Pred)
12+
open import Function.Bundles using (_⟶_; _↣_; _↠_; _⤖_; _⇔_; _↩_; _↪_; _↩↪_; _↔_)
13+
open import Relation.Binary hiding (_⇔_)
14+
open import Level using (Level)
15+
16+
private
17+
variable
18+
a b ℓ₁ ℓ₂ : Level
19+
20+
------------------------------------------------------------------------
21+
-- Bundles specialised for lifting relations to indexed sets
22+
------------------------------------------------------------------------
23+
24+
infix 3 _⟶ᵢ_ _↣ᵢ_ _↠ᵢ_ _⤖ᵢ_ _⇔ᵢ_ _↩ᵢ_ _↪ᵢ_ _↩↪ᵢ_ _↔ᵢ_
25+
_⟶ᵢ_ : {i} {I : Set i} Pred I a Pred I b Set _
26+
A ⟶ᵢ B = {i} A i ⟶ B i
27+
28+
_↣ᵢ_ : {i} {I : Set i} Pred I a Pred I b Set _
29+
A ↣ᵢ B = {i} A i ↣ B i
30+
31+
_↠ᵢ_ : {i} {I : Set i} Pred I a Pred I b Set _
32+
A ↠ᵢ B = {i} A i ↠ B i
33+
34+
_⤖ᵢ_ : {i} {I : Set i} Pred I a Pred I b Set _
35+
A ⤖ᵢ B = {i} A i ⤖ B i
36+
37+
_⇔ᵢ_ : {i} {I : Set i} Pred I a Pred I b Set _
38+
A ⇔ᵢ B = {i} A i ⇔ B i
39+
40+
_↩ᵢ_ : {i} {I : Set i} Pred I a Pred I b Set _
41+
A ↩ᵢ B = {i} A i ↩ B i
42+
43+
_↪ᵢ_ : {i} {I : Set i} Pred I a Pred I b Set _
44+
A ↪ᵢ B = {i} A i ↪ B i
45+
46+
_↩↪ᵢ_ : {i} {I : Set i} Pred I a Pred I b Set _
47+
A ↩↪ᵢ B = {i} A i ↩↪ B i
48+
49+
_↔ᵢ_ : {i} {I : Set i} Pred I a Pred I b Set _
50+
A ↔ᵢ B = {i} A i ↔ B i

0 commit comments

Comments
 (0)