Skip to content

Commit cc7d32d

Browse files
guilhermehasguialvaresMatthewDaggitt
authored
Added tabulate+< (#2190)
* added tabulate+< * renamed tabulate function --------- Co-authored-by: Guilherme <[email protected]> Co-authored-by: MatthewDaggitt <[email protected]>
1 parent 4437fcb commit cc7d32d

File tree

2 files changed

+16
-7
lines changed

2 files changed

+16
-7
lines changed

CHANGELOG.md

Lines changed: 5 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -26,3 +26,8 @@ New modules
2626

2727
Additions to existing modules
2828
-----------------------------
29+
30+
* In `Data.List.Relation.Unary.AllPairs.Properties`:
31+
```
32+
tabulate⁺-< : (i < j → R (f i) (f j)) → AllPairs R (tabulate f)
33+
```

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

Lines changed: 11 additions & 7 deletions
Original file line numberDiff line numberDiff line change
@@ -13,9 +13,9 @@ open import Data.List.Relation.Unary.All as All using (All; []; _∷_)
1313
import Data.List.Relation.Unary.All.Properties as All
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.Fin.Base using (Fin)
17-
open import Data.Fin.Properties using (suc-injective)
18-
open import Data.Nat.Base using (zero; suc; _<_; z≤n; s≤s)
16+
open import Data.Fin.Base as F using (Fin)
17+
open import Data.Fin.Properties using (suc-injective; <⇒≢)
18+
open import Data.Nat.Base using (zero; suc; _<_; z≤n; s≤s; z<s; s<s)
1919
open import Data.Nat.Properties using (≤-refl; m<n⇒m<1+n)
2020
open import Function.Base using (_∘_; flip)
2121
open import Level using (Level)
@@ -115,12 +115,16 @@ module _ {R : Rel A ℓ} where
115115

116116
module _ {R : Rel A ℓ} where
117117

118+
tabulate⁺-< : {n} {f : Fin n A} ( {i j} i F.< j R (f i) (f j))
119+
AllPairs R (tabulate f)
120+
tabulate⁺-< {zero} fᵢ~fⱼ = []
121+
tabulate⁺-< {suc n} fᵢ~fⱼ =
122+
All.tabulate⁺ (λ _ fᵢ~fⱼ z<s) ∷
123+
tabulate⁺-< (fᵢ~fⱼ ∘ s<s)
124+
118125
tabulate⁺ : {n} {f : Fin n A} ( {i j} i ≢ j R (f i) (f j))
119126
AllPairs R (tabulate f)
120-
tabulate⁺ {zero} fᵢ~fⱼ = []
121-
tabulate⁺ {suc n} fᵢ~fⱼ =
122-
All.tabulate⁺ (λ j fᵢ~fⱼ λ()) ∷
123-
tabulate⁺ (fᵢ~fⱼ ∘ (_∘ suc-injective))
127+
tabulate⁺ fᵢ~fⱼ = tabulate⁺-< (fᵢ~fⱼ ∘ <⇒≢)
124128

125129
------------------------------------------------------------------------
126130
-- filter

0 commit comments

Comments
 (0)