@@ -13,9 +13,9 @@ open import Data.List.Relation.Unary.All as All using (All; []; _∷_)
1313import Data.List.Relation.Unary.All.Properties as All
1414open import Data.List.Relation.Unary.AllPairs as AllPairs using (AllPairs; []; _∷_)
1515open 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 )
1919open import Data.Nat.Properties using (≤-refl; m<n⇒m<1+n)
2020open import Function.Base using (_∘_; flip)
2121open import Level using (Level)
@@ -115,12 +115,16 @@ module _ {R : Rel A ℓ} where
115115
116116module _ {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