diff --git a/CHANGELOG.md b/CHANGELOG.md index 64b9eb8db1..155d93d480 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -335,6 +335,28 @@ Deprecated names sym-↔ ↦ ↔-sym ``` +* In `Data.Fin.Base`: +two new, hopefully more memorable, names `↑ˡ` `↑ʳ` for the 'left', resp. 'right' injection of a Fin m into a 'larger' type, `Fin (m + n)`, resp. `Fin (n + m)`, with argument order to reflect the position of the Fin m argument. + ``` + inject+ ↦ flip _↑ˡ_ + raise ↦ _↑ʳ_ + ``` + +* In `Data.Fin.Properties`: + ``` + toℕ-raise ↦ toℕ-↑ʳ + toℕ-inject+ n i ↦ sym (toℕ-↑ˡ i n) + splitAt-inject+ m n i ↦ splitAt-↑ˡ m i n + splitAt-raise ↦ splitAt-↑ʳ + ``` + +* In `Data.Vec.Properties`: + ``` + []≔-++-inject+ ↦ []≔-++-↑ˡ + ``` + Additionally, `[]≔-++-↑ʳ`, by analogy. + + New modules ----------- @@ -506,6 +528,7 @@ Other minor changes map-⊛ : ∀ {n} (f : A → B → C) (g : A → B) (xs : Vec A n) → (map f xs ⊛ map g xs) ≡ map (f ˢ g) xs ⊛-is->>= : ∀ {n} (fs : Vec (A → B) n) (xs : Vec A n) → (fs ⊛ xs) ≡ (fs DiagonalBind.>>= flip map xs) transpose-replicate : ∀ {m n} (xs : Vec A m) → transpose (replicate {n = n} xs) ≡ map replicate xs + []≔-++-↑ʳ : ∀ {m n y} (xs : Vec A m) (ys : Vec A n) i → (xs ++ ys) [ m ↑ʳ i ]≔ y ≡ xs ++ (ys [ i ]≔ y) ``` * Added new proofs in `Function.Construct.Symmetry`: diff --git a/HACKING.md b/HACKING.md index 448ef3ae43..9b59aa6516 100644 --- a/HACKING.md +++ b/HACKING.md @@ -81,7 +81,7 @@ How to make changes thrown up until the tests are passed. Note that the tests require the use of a tool called `fix-whitespace`. See the instructions at the end of this file for how to install this. - + Note this step is optional as these tests will also be run automatically by our CI infrastructure when you open a pull request on Github, but it can be useful to run it locally to get a faster turn around time when finding diff --git a/src/Data/Fin/Base.agda b/src/Data/Fin/Base.agda index 54c76e9b61..7186d73bb6 100644 --- a/src/Data/Fin/Base.agda +++ b/src/Data/Fin/Base.agda @@ -17,7 +17,7 @@ open import Data.Nat.Base as ℕ using (ℕ; zero; suc; z≤n; s≤s) open import Data.Nat.Properties.Core using (≤-pred) open import Data.Product as Product using (_×_; _,_) open import Data.Sum.Base as Sum using (_⊎_; inj₁; inj₂; [_,_]′) -open import Function.Base using (id; _∘_; _on_) +open import Function.Base using (id; _∘_; _on_; flip) open import Level using (0ℓ) open import Relation.Nullary using (yes; no) open import Relation.Nullary.Decidable.Core using (True; toWitness) @@ -77,11 +77,19 @@ fromℕ<″ zero (ℕ.less-than-or-equal refl) = zero fromℕ<″ (suc m) (ℕ.less-than-or-equal refl) = suc (fromℕ<″ m (ℕ.less-than-or-equal refl)) --- raise m "i" = "m + i". +-- canonical liftings of i:Fin m to larger index -raise : ∀ {m} n → Fin m → Fin (n ℕ.+ m) -raise zero i = i -raise (suc n) i = suc (raise n i) +-- injection on the left: "i" ↑ˡ n = "i" in Fin (m + n) +infixl 5 _↑ˡ_ +_↑ˡ_ : ∀ {m} → Fin m → ∀ n → Fin (m ℕ.+ n) +zero ↑ˡ n = zero +(suc i) ↑ˡ n = suc (i ↑ˡ n) + +-- injection on the right: n ↑ʳ "i" = "n + i" in Fin (n + m) +infixr 5 _↑ʳ_ +_↑ʳ_ : ∀ {m} n → Fin m → Fin (n ℕ.+ m) +zero ↑ʳ i = i +(suc n) ↑ʳ i = suc (n ↑ʳ i) -- reduce≥ "m + i" _ = "i". @@ -99,10 +107,6 @@ inject! : ∀ {n} {i : Fin (suc n)} → Fin′ i → Fin n inject! {n = suc _} {i = suc _} zero = zero inject! {n = suc _} {i = suc _} (suc j) = suc (inject! j) -inject+ : ∀ {m} n → Fin m → Fin (m ℕ.+ n) -inject+ n zero = zero -inject+ n (suc i) = suc (inject+ n i) - inject₁ : ∀ {m} → Fin m → Fin (suc m) inject₁ zero = zero inject₁ (suc i) = suc (inject₁ i) @@ -134,7 +138,7 @@ splitAt (suc m) (suc i) = Sum.map suc id (splitAt m i) -- inverse of above function join : ∀ m n → Fin m ⊎ Fin n → Fin (m ℕ.+ n) -join m n = [ inject+ n , raise {n} m ]′ +join m n = [ _↑ˡ n , m ↑ʳ_ ]′ -- quotRem k "i" = "i % k" , "i / k" -- This is dual to group from Data.Vec. @@ -150,8 +154,8 @@ remQuot k = Product.swap ∘ quotRem k -- inverse of remQuot combine : ∀ {n k} → Fin n → Fin k → Fin (n ℕ.* k) -combine {suc n} {k} zero y = inject+ (n ℕ.* k) y -combine {suc n} {k} (suc x) y = raise k (combine x y) +combine {suc n} {k} zero y = y ↑ˡ (n ℕ.* k) +combine {suc n} {k} (suc x) y = k ↑ʳ (combine x y) ------------------------------------------------------------------------ -- Operations @@ -307,3 +311,20 @@ fromℕ≤″ = fromℕ<″ "Warning: fromℕ≤″ was deprecated in v1.2. Please use fromℕ<″ instead." #-} + +-- Version 2.0 + +raise = _↑ʳ_ +{-# WARNING_ON_USAGE raise +"Warning: raise was deprecated in v2.0. +Please use _↑_ʳ instead." +#-} +inject+ : ∀ {m} n → Fin m → Fin (m ℕ.+ n) +inject+ n i = i ↑ˡ n +{-# WARNING_ON_USAGE inject+ +"Warning: inject+ was deprecated in v2.0. +Please use _↑ˡ_ instead. +NB argument order has been flipped: +the left-hand argument is the Fin m +the right-hand is the Nat index increment." +#-} diff --git a/src/Data/Fin/Properties.agda b/src/Data/Fin/Properties.agda index 317e349d0d..c6d3a19aea 100644 --- a/src/Data/Fin/Properties.agda +++ b/src/Data/Fin/Properties.agda @@ -21,7 +21,7 @@ open import Data.Unit using (tt) open import Data.Product using (Σ-syntax; ∃; ∃₂; ∄; _×_; _,_; map; proj₁; uncurry; <_,_>) open import Data.Sum.Base as Sum using (_⊎_; inj₁; inj₂; [_,_]; [_,_]′) open import Data.Sum.Properties using ([,]-map-commute; [,]-∘-distr) -open import Function.Base using (_∘_; id; _$_) +open import Function.Base using (_∘_; id; _$_; flip) open import Function.Bundles using (_↔_; mk↔′) open import Function.Definitions.Core2 using (Surjective) open import Function.Equivalence using (_⇔_; equivalence) @@ -105,9 +105,21 @@ toℕ-strengthen : ∀ {n} (i : Fin n) → toℕ (strengthen i) ≡ toℕ i toℕ-strengthen zero = refl toℕ-strengthen (suc i) = cong suc (toℕ-strengthen i) -toℕ-raise : ∀ {m} n (i : Fin m) → toℕ (raise n i) ≡ n ℕ.+ toℕ i -toℕ-raise zero i = refl -toℕ-raise (suc n) i = cong suc (toℕ-raise n i) +------------------------------------------------------------------------ +-- toℕ-↑ˡ: "i" ↑ˡ n = "i" in Fin (m + n) +------------------------------------------------------------------------ + +toℕ-↑ˡ : ∀ {m} (i : Fin m) n → toℕ (i ↑ˡ n) ≡ toℕ i +toℕ-↑ˡ zero n = refl +toℕ-↑ˡ (suc i) n = cong suc (toℕ-↑ˡ i n) + +------------------------------------------------------------------------ +-- toℕ-↑ʳ: n ↑ʳ "i" = "n + i" in Fin (n + m) +------------------------------------------------------------------------ + +toℕ-↑ʳ : ∀ {m} n (i : Fin m) → toℕ (n ↑ʳ i) ≡ n ℕ.+ toℕ i +toℕ-↑ʳ zero i = refl +toℕ-↑ʳ (suc n) i = cong suc (toℕ-↑ʳ n i) toℕ