From 535e4297a08c626d0e2e1923914727f914b8c9bd Mon Sep 17 00:00:00 2001 From: Greg Brown Date: Sat, 19 Mar 2022 13:28:42 +0000 Subject: Modify pseudocode definition. This change makes the following changes to the definition of pseudocode: - Add a separate type `bit` for single-bit values. - Change `var` and `state` to take `Fin`s instead of bounded naturals. - Make `[_]` and `unbox` work for any sliced type. - Generalise `_:_` and `slice` into `splice` and `cut` respectively, making the two new operations inverses. - Replace `tup` with `nil` and `cons` for building tuples. - Add destructors `head` and `tail` for tuple types. - Make function and procedure calls take a vector of arguments instead of a tuple. - Introduce an `if_then_` expression for if statements with a trivial else branch. --- src/Helium/Data/Pseudocode/Core.agda | 261 ++++++++++++++----------- src/Helium/Data/Pseudocode/Types.agda | 2 + src/Helium/Instructions/Base.agda | 145 +++++++------- src/Helium/Instructions/Instances/Barrett.agda | 4 +- src/Helium/Semantics/Denotational/Core.agda | 156 +++++++++------ 5 files changed, 319 insertions(+), 249 deletions(-) (limited to 'src') diff --git a/src/Helium/Data/Pseudocode/Core.agda b/src/Helium/Data/Pseudocode/Core.agda index 8a0c4e3..079e2ce 100644 --- a/src/Helium/Data/Pseudocode/Core.agda +++ b/src/Helium/Data/Pseudocode/Core.agda @@ -9,7 +9,8 @@ module Helium.Data.Pseudocode.Core where open import Data.Bool using (Bool; true; false) -open import Data.Fin using (Fin; #_) +open import Data.Fin as Fin using (Fin) +open import Data.Fin.Patterns open import Data.Nat as ℕ using (ℕ; zero; suc) open import Data.Nat.Properties using (+-comm) open import Data.Product using (∃; _,_; proj₂; uncurry) @@ -19,7 +20,7 @@ open import Data.Vec.Relation.Unary.All using (All; []; _∷_; reduce) open import Data.Vec.Relation.Unary.Any using (Any; here; there) open import Function as F using (_∘_; _∘′_; _∋_) open import Relation.Binary.PropositionalEquality using (_≡_; refl) -open import Relation.Nullary using (yes; no) +open import Relation.Nullary using (Dec; yes; no) open import Relation.Nullary.Decidable.Core using (True; False; toWitness; fromWitness; map′) open import Relation.Nullary.Product using (_×-dec_) open import Relation.Nullary.Sum using (_⊎-dec_) @@ -31,18 +32,17 @@ data Type : Set where int : Type fin : (n : ℕ) → Type real : Type + bit : Type bits : (n : ℕ) → Type tuple : ∀ n → Vec Type n → Type array : Type → (n : ℕ) → Type -bit : Type -bit = bits 1 - data HasEquality : Type → Set where bool : HasEquality bool int : HasEquality int fin : ∀ {n} → HasEquality (fin n) real : HasEquality real + bit : HasEquality bit bits : ∀ {n} → HasEquality (bits n) hasEquality? : Decidable HasEquality @@ -50,6 +50,7 @@ hasEquality? bool = yes bool hasEquality? int = yes int hasEquality? (fin n) = yes fin hasEquality? real = yes real +hasEquality? bit = yes bit hasEquality? (bits n) = yes bits hasEquality? (tuple n x) = no (λ ()) hasEquality? (array t n) = no (λ ()) @@ -61,16 +62,17 @@ data IsNumeric : Type → Set where isNumeric? : Decidable IsNumeric isNumeric? bool = no (λ ()) isNumeric? int = yes int -isNumeric? real = yes real isNumeric? (fin n) = no (λ ()) +isNumeric? real = yes real +isNumeric? bit = no (λ ()) isNumeric? (bits n) = no (λ ()) isNumeric? (tuple n x) = no (λ ()) isNumeric? (array t n) = no (λ ()) -combineNumeric : ∀ t₁ t₂ → {isNumeric₁ : True (isNumeric? t₁)} → {isNumeric₂ : True (isNumeric? t₂)} → Type -combineNumeric int int = int -combineNumeric int real = real -combineNumeric real _ = real +combineNumeric : ∀ t₁ t₂ → (isNumeric₁ : True (isNumeric? t₁)) → (isNumeric₂ : True (isNumeric? t₂)) → Type +combineNumeric int int _ _ = int +combineNumeric int real _ _ = real +combineNumeric real _ _ _ = real data Sliced : Set where bits : Sliced @@ -87,12 +89,13 @@ elemType (array t) = t --- Literals data Literal : Type → Set where - _′b : Bool → Literal bool - _′i : ℕ → Literal int - _′f : ∀ {n} → Fin n → Literal (fin n) - _′r : ℕ → Literal real - _′x : ∀ {n} → Vec Bool n → Literal (bits n) - _′a : ∀ {n t} → Literal t → Literal (array t n) + _′b : Bool → Literal bool + _′i : ℕ → Literal int + _′f : ∀ {n} → Fin n → Literal (fin n) + _′r : ℕ → Literal real + _′x : Bool → Literal bit + _′xs : ∀ {n} → Vec Bool n → Literal (bits n) + _′a : ∀ {n t} → Literal t → Literal (array t n) --- Expressions, references, statements, functions and procedures @@ -100,10 +103,8 @@ module Code {o} (Σ : Vec Type o) where data Expression {n} (Γ : Vec Type n) : Type → Set data CanAssign {n Γ} : ∀ {t} → Expression {n} Γ t → Set canAssign? : ∀ {n Γ t} → Decidable (CanAssign {n} {Γ} {t}) - canAssignAll? : ∀ {n Γ m ts} → Decidable {A = All (Expression {n} Γ) {m} ts} (All (CanAssign ∘ proj₂) ∘ (reduce (λ {x} e → x , e))) - data AssignsState {n Γ} : ∀ {t e} → CanAssign {n} {Γ} {t} e → Set - assignsState? : ∀ {n Γ t e} → Decidable (AssignsState {n} {Γ} {t} {e}) - assignsStateAny? : ∀ {n Γ m ts es} → Decidable {A = All (CanAssign ∘ proj₂) (reduce {P = Expression {n} Γ} (λ {x} e → x , e) {m} {ts} es)} (Any (AssignsState ∘ proj₂) ∘ reduce (λ {x} a → x , a)) + data ReferencesState {n Γ} : ∀ {t} → Expression {n} Γ t → Set + referencesState? : ∀ {n Γ t} → Decidable (ReferencesState {n} {Γ} {t}) data Statement {n} (Γ : Vec Type n) : Set data ChangesState {n Γ} : Statement {n} Γ → Set changesState? : ∀ {n Γ} → Decidable (ChangesState {n} {Γ}) @@ -114,55 +115,61 @@ module Code {o} (Σ : Vec Type o) where infixr 7 _^_ infixl 6 _*_ _and_ _>>_ -- infixl 6 _/_ - infixl 5 _+_ _or_ _&&_ _||_ _∶_ + infixl 5 _+_ _or_ _&&_ _||_ infix 4 _≟_ _>_ : Expression Γ int → ℕ → Expression Γ int - rnd : Expression Γ real → Expression Γ int - fin : ∀ {k ms n} → (All (Fin) ms → Fin n) → Expression Γ (tuple k (map fin ms)) → Expression Γ (fin n) - asInt : ∀ {m} → Expression Γ (fin m) → Expression Γ int - tup : ∀ {m ts} → All (Expression Γ) ts → Expression Γ (tuple m ts) - call : ∀ {t m Δ} → Function Δ t → Expression Γ (tuple m Δ) → Expression Γ t + _^_ : ∀ {t} {isNumeric : True (isNumeric? t)} → Expression Γ t → ℕ → Expression Γ t + _>>_ : Expression Γ int → ℕ → Expression Γ int + rnd : Expression Γ real → Expression Γ int + fin : ∀ {k ms n} → (All (Fin) ms → Fin n) → Expression Γ (tuple k (map fin ms)) → Expression Γ (fin n) + asInt : ∀ {m} → Expression Γ (fin m) → Expression Γ int + nil : Expression Γ (tuple 0 []) + cons : ∀ {m t ts} → Expression Γ t → Expression Γ (tuple m ts) → Expression Γ (tuple (suc m) (t ∷ ts)) + head : ∀ {m t ts} → Expression Γ (tuple (suc m) (t ∷ ts)) → Expression Γ t + tail : ∀ {m t ts} → Expression Γ (tuple (suc m) (t ∷ ts)) → Expression Γ (tuple m ts) + call : ∀ {t m Δ} → Function Δ t → All (Expression Γ) {m} Δ → Expression Γ t if_then_else_ : ∀ {t} → Expression Γ bool → Expression Γ t → Expression Γ t → Expression Γ t data CanAssign {n} {Γ} where - state : ∀ i {i> x) = no λ () + referencesState? (rnd e) = no λ () + referencesState? (fin x e) = no λ () + referencesState? (asInt e) = no λ () + referencesState? nil = no λ () + referencesState? (cons e e₁) = map′ [ (λ e → consˡ e e₁) , (λ e₁ → consʳ e e₁) ]′ (λ { (consˡ e e₁) → inj₁ e ; (consʳ e e₁) → inj₂ e₁ }) (referencesState? e ⊎-dec referencesState? e₁) + referencesState? (head e) = map′ head (λ { (head e) → e }) (referencesState? e) + referencesState? (tail e) = map′ tail (λ { (tail e) → e }) (referencesState? e) + referencesState? (call f es) = no λ () + referencesState? (if e then e₁ else e₂) = no λ () infix 4 _≔_ - infixl 2 if_then_else_ + infixl 2 if_then_else_ if_then_ infixl 1 _∙_ _∙return_ infix 1 _∙end @@ -225,27 +254,30 @@ module Code {o} (Σ : Vec Type o) where skip : Statement Γ _≔_ : ∀ {t} → (ref : Expression Γ t) → {canAssign : True (canAssign? ref)} → Expression Γ t → Statement Γ declare : ∀ {t} → Expression Γ t → Statement (t ∷ Γ) → Statement Γ - invoke : ∀ {m Δ} → Procedure Δ → Expression Γ (tuple m Δ) → Statement Γ + invoke : ∀ {m Δ} → Procedure Δ → All (Expression Γ) {m} Δ → Statement Γ + if_then_ : Expression Γ bool → Statement Γ → Statement Γ if_then_else_ : Expression Γ bool → Statement Γ → Statement Γ → Statement Γ for : ∀ m → Statement (fin m ∷ Γ) → Statement Γ data ChangesState where - _∙ˡ_ : ∀ {s₁} → ChangesState s₁ → ∀ s₂ → ChangesState (s₁ ∙ s₂) - _∙ʳ_ : ∀ s₁ {s₂} → ChangesState s₂ → ChangesState (s₁ ∙ s₂) - _≔_ : ∀ {t ref} canAssign {assignsState : True (assignsState? (toWitness canAssign))} e₂ → ChangesState (_≔_ {t = t} ref {canAssign} e₂) + _∙ˡ_ : ∀ {s} → ChangesState s → ∀ s₁ → ChangesState (s ∙ s₁) + _∙ʳ_ : ∀ s {s₁} → ChangesState s₁ → ChangesState (s ∙ s₁) + _≔_ : ∀ {t} ref {canAssign : True (canAssign? ref)} {refsState : True (referencesState? ref)} e₂ → ChangesState (_≔_ {t = t} ref {canAssign} e₂) declare : ∀ {t} e {s} → ChangesState s → ChangesState (declare {t = t} e s) - invoke : ∀ {m Δ} p e → ChangesState (invoke {m = m} {Δ} p e) - if_then′_else_ : ∀ e {s₁} → ChangesState s₁ → ∀ s₂ → ChangesState (if e then s₁ else s₂) - if_then_else′_ : ∀ e s₁ {s₂} → ChangesState s₂ → ChangesState (if e then s₁ else s₂) + invoke : ∀ {m Δ} p es → ChangesState (invoke {m = m} {Δ} p es) + if_then_ : ∀ e {s} → ChangesState s → ChangesState (if e then s) + if_then′_else_ : ∀ e {s} → ChangesState s → ∀ s₁ → ChangesState (if e then s else s₁) + if_then_else′_ : ∀ e s {s₁} → ChangesState s₁ → ChangesState (if e then s else s₁) for : ∀ m {s} → ChangesState s → ChangesState (for m s) - changesState? (s₁ ∙ s₂) = map′ [ _∙ˡ s₂ , s₁ ∙ʳ_ ]′ (λ { (s ∙ˡ _) → inj₁ s ; (_ ∙ʳ s) → inj₂ s }) (changesState? s₁ ⊎-dec changesState? s₂) - changesState? skip = no λ () - changesState? (_≔_ ref {a} e) = map′ (λ s → _≔_ a {fromWitness s} e) (λ { (_≔_ _ {s} _) → toWitness s }) (assignsState? (toWitness a)) - changesState? (declare e s) = map′ (declare e) (λ { (declare e s) → s }) (changesState? s) - changesState? (invoke p e) = yes (invoke p e) - changesState? (if e then s₁ else s₂) = map′ [ if e then′_else s₂ , if e then s₁ else′_ ]′ (λ { (if _ then′ s else _) → inj₁ s ; (if _ then _ else′ s) → inj₂ s }) (changesState? s₁ ⊎-dec changesState? s₂) - changesState? (for m s) = map′ (for m) (λ { (for m s) → s }) (changesState? s) + changesState? (s ∙ s₁) = map′ [ _∙ˡ s₁ , s ∙ʳ_ ]′ (λ { (s ∙ˡ s₁) → inj₁ s ; (s ∙ʳ s₁) → inj₂ s₁ }) (changesState? s ⊎-dec changesState? s₁) + changesState? skip = no λ () + changesState? (_≔_ ref e) = map′ (λ refsState → _≔_ ref {refsState = fromWitness refsState} e) (λ { (_≔_ ref {refsState = refsState} e) → toWitness refsState }) (referencesState? ref) + changesState? (declare e s) = map′ (declare e) (λ { (declare e s) → s }) (changesState? s) + changesState? (invoke p e) = yes (invoke p e) + changesState? (if e then s) = map′ (if e then_) (λ { (if e then s) → s }) (changesState? s) + changesState? (if e then s else s₁) = map′ [ if e then′_else s₁ , if e then s else′_ ]′ (λ { (if e then′ s else s₁) → inj₁ s ; (if e then s else′ s₁) → inj₂ s₁ }) (changesState? s ⊎-dec changesState? s₁) + changesState? (for m s) = map′ (for m) (λ { (for m s) → s }) (changesState? s) data Function Γ ret where _∙return_ : (s : Statement Γ) → {False (changesState? s)} → Expression Γ ret → Function Γ ret @@ -253,36 +285,45 @@ module Code {o} (Σ : Vec Type o) where data Procedure Γ where _∙end : Statement Γ → Procedure Γ - declare : ∀ {t} → Expression Γ t → Procedure (t ∷ Γ) → Procedure Γ infixl 6 _<<_ - infixl 5 _-_ + infixl 5 _-_ _∶_ + + tup : ∀ {n Γ m ts} → All (Expression {n} Γ) ts → Expression Γ (tuple m ts) + tup [] = nil + tup (e ∷ es) = cons e (tup es) + + _∶_ : ∀ {n Γ i j t} → Expression {n} Γ (asType t j) → Expression Γ (asType t i) → Expression Γ (asType t (i ℕ.+ j)) + e₁ ∶ e₂ = splice e₁ e₂ (lit (Fin.fromℕ _ ′f)) + + slice : ∀ {n Γ i j t} → Expression {n} Γ (asType t (i ℕ.+ j)) → Expression Γ (fin (suc i)) → Expression Γ (asType t j) + slice e₁ e₂ = head (cut e₁ e₂) slice′ : ∀ {n Γ i j t} → Expression {n} Γ (asType t (i ℕ.+ j)) → Expression Γ (fin (suc j)) → Expression Γ (asType t i) slice′ {i = i} e₁ e₂ = slice (cast (+-comm i _) e₁) e₂ - _-_ : ∀ {n Γ t₁ t₂} {isNumeric₁ : True (isNumeric? t₁)} {isNumeric₂ : True (isNumeric? t₂)} → Expression {n} Γ t₁ → Expression Γ t₂ → Expression Γ (combineNumeric t₁ t₂ {isNumeric₁} {isNumeric₂}) + _-_ : ∀ {n Γ t₁ t₂} {isNumeric₁ : True (isNumeric? t₁)} {isNumeric₂ : True (isNumeric? t₂)} → Expression {n} Γ t₁ → Expression Γ t₂ → Expression Γ (combineNumeric t₁ t₂ isNumeric₁ isNumeric₂) _-_ {isNumeric₂ = isNumeric₂} x y = x + (-_ {isNumeric = isNumeric₂} y) _<<_ : ∀ {n Γ} → Expression {n} Γ int → ℕ → Expression Γ int x << n = rnd (x * lit (2 ′r) ^ n) get : ∀ {n Γ} → ℕ → Expression {n} Γ int → Expression Γ bit - get i x = if x - x >> suc i << suc i > suc i << suc i >3 (tup ((var 3) ∷ []))) + tup (index (var 5F) (var 3F) ∷ var 0F ∷ []) ≔ call (SignedSatQ (toℕ esize-1)) (var 1F ∷ []) ∙ + if var 0F && hasBit (var 6F) (fin e*esize>>3 (tup ((var 3F) ∷ []))) then - FPSCR-QC ≔ lit ((true ∷ []) ′x) - else skip))) + FPSCR-QC ≔ lit (true ′x)))) where open Instr.VecOp₂ d @@ -299,9 +292,9 @@ private helper Instr.32bit i = Fin.combine i zero vqdmulh : Instr.VQDMulH → Procedure [] -vqdmulh d = vqr?dmulh d (skip ∙return lit (2 ′i) * var 0 * var 1 >> toℕ esize) +vqdmulh d = vqr?dmulh d (skip ∙return lit (2 ′i) * var 0F * var 1F >> toℕ esize) where open Instr.VecOp₂ d using (esize) vqrdmulh : Instr.VQRDMulH → Procedure [] -vqrdmulh d = vqr?dmulh d (skip ∙return lit (2 ′i) * var 0 * var 1 + lit (1 ′i) << toℕ esize-1 >> toℕ esize) +vqrdmulh d = vqr?dmulh d (skip ∙return lit (2 ′i) * var 0F * var 1F + lit (1 ′i) << toℕ esize-1 >> toℕ esize) where open Instr.VecOp₂ d using (esize; esize-1) diff --git a/src/Helium/Instructions/Instances/Barrett.agda b/src/Helium/Instructions/Instances/Barrett.agda index 5ec9ba4..606a9e9 100644 --- a/src/Helium/Instructions/Instances/Barrett.agda +++ b/src/Helium/Instructions/Instances/Barrett.agda @@ -25,9 +25,9 @@ open import Helium.Instructions.Core barret : (m -n : Expression [] (bits 32)) (t z : VecReg) (im : GenReg) → Procedure [] barret m -n t z im = index R (lit (im ′f)) ≔ m ∙ - invoke vqrdmulh-s32,t,z,m (tup []) ∙ + invoke vqrdmulh-s32,t,z,m [] ∙ index R (lit (im ′f)) ≔ -n ∙ - invoke vmla-s32,z,t,-n (tup []) ∙end + invoke vmla-s32,z,t,-n [] ∙end where vqrdmulh-s32,t,z,m = ExecBeats (vqrdmulh (record diff --git a/src/Helium/Semantics/Denotational/Core.agda b/src/Helium/Semantics/Denotational/Core.agda index b425252..0bd1794 100644 --- a/src/Helium/Semantics/Denotational/Core.agda +++ b/src/Helium/Semantics/Denotational/Core.agda @@ -17,7 +17,7 @@ private open module C = RawPseudocode rawPseudocode open import Data.Bool as Bool using (Bool; true; false) -open import Data.Fin as Fin using (Fin; zero; suc; #_) +open import Data.Fin as Fin using (Fin; zero; suc) import Data.Fin.Properties as Finₚ open import Data.Nat as ℕ using (ℕ; zero; suc) import Data.Nat.Properties as ℕₚ @@ -31,7 +31,7 @@ open import Function using (case_of_; _∘′_; id) open import Helium.Data.Pseudocode.Core import Induction as I import Induction.WellFounded as Wf -open import Level hiding (zero; suc) +open import Level using (Level; _⊔_; 0ℓ) open import Relation.Binary.PropositionalEquality as ≡ using (_≡_; module ≡-Reasoning) open import Relation.Nullary using (does) open import Relation.Nullary.Decidable.Core using (True; False; toWitness; fromWitness) @@ -41,6 +41,7 @@ open import Relation.Nullary.Decidable.Core using (True; False; toWitness; fromW ⟦ int ⟧ₗ = i₁ ⟦ fin n ⟧ₗ = 0ℓ ⟦ real ⟧ₗ = r₁ +⟦ bit ⟧ₗ = b₁ ⟦ bits n ⟧ₗ = b₁ ⟦ tuple n ts ⟧ₗ = helper ts where @@ -56,6 +57,7 @@ open import Relation.Nullary.Decidable.Core using (True; False; toWitness; fromW ⟦ int ⟧ₜ = ℤ ⟦ fin n ⟧ₜ = Fin n ⟦ real ⟧ₜ = ℝ +⟦ bit ⟧ₜ = Bit ⟦ bits n ⟧ₜ = Bits n ⟦ tuple n ts ⟧ₜ = ⟦ ts ⟧ₜ′ ⟦ array t n ⟧ₜ = Vec ⟦ t ⟧ₜ n @@ -66,13 +68,13 @@ open import Relation.Nullary.Decidable.Core using (True; False; toWitness; fromW -- The case for bitvectors looks odd so that the right-most bit is bit 0. 𝒦 : ∀ {t} → Literal t → ⟦ t ⟧ₜ -𝒦 (x ′b) = x -𝒦 (x ′i) = x ℤ′.×′ 1ℤ -𝒦 (x ′f) = x -𝒦 (x ′r) = x ℝ′.×′ 1ℝ -𝒦 (xs ′x) = Vec.foldl Bits (λ bs b → (Bool.if b then 1𝔹 else 0𝔹) VecF.∷ bs) VecF.[] xs -𝒦 (x ′a) = Vec.replicate (𝒦 x) - +𝒦 (x ′b) = x +𝒦 (x ′i) = x ℤ′.×′ 1ℤ +𝒦 (x ′f) = x +𝒦 (x ′r) = x ℝ′.×′ 1ℝ +𝒦 (x ′x) = Bool.if x then 1𝔹 else 0𝔹 +𝒦 (xs ′xs) = Vec.foldl Bits (λ bs b → (Bool.if b then 1𝔹 else 0𝔹) VecF.∷ bs) VecF.[] xs +𝒦 (x ′a) = Vec.replicate (𝒦 x) fetch : ∀ {n} ts → ⟦ tuple n ts ⟧ₜ → ∀ i → ⟦ Vec.lookup ts i ⟧ₜ fetch (_ ∷ []) x zero = x @@ -96,10 +98,11 @@ tupTail [] x = _ tupTail (_ ∷ _) (x , xs) = xs equal : ∀ {t} → HasEquality t → ⟦ t ⟧ₜ → ⟦ t ⟧ₜ → Bool -equal bool x y = does (x Bool.≟ y) -equal int x y = does (x ≟ᶻ y) -equal fin x y = does (x Fin.≟ y) -equal real x y = does (x ≟ʳ y) +equal bool x y = does (x Bool.≟ y) +equal int x y = does (x ≟ᶻ y) +equal fin x y = does (x Fin.≟ y) +equal real x y = does (x ≟ʳ y) +equal bit x y = does (x ≟ᵇ₁ y) equal bits x y = does (x ≟ᵇ y) comp : ∀ {t} → IsNumeric t → ⟦ t ⟧ₜ → ⟦ t ⟧ₜ → Bool @@ -108,8 +111,8 @@ comp real x y = does (x <ʳ? y) -- 0 of y is 0 of result join : ∀ t {m n} → ⟦ asType t m ⟧ₜ → ⟦ asType t n ⟧ₜ → ⟦ asType t (n ℕ.+ m) ⟧ₜ -join bits {m} x y = y VecF.++ x -join (array _) {m} x y = y Vec.++ x +join bits x y = y VecF.++ x +join (array _) x y = y Vec.++ x -- take from 0 take : ∀ t i {j} → ⟦ asType t (i ℕ.+ j) ⟧ₜ → ⟦ asType t i ⟧ₜ @@ -131,38 +134,59 @@ module _ where m≤n⇒m+k≡n ℕ.z≤n = _ , ≡.refl m≤n⇒m+k≡n (ℕ.s≤s m≤n) = P.dmap id (≡.cong suc) (m≤n⇒m+k≡n m≤n) - slicedSize : ∀ i j (off : Fin (suc i)) → P.∃ λ k → i ℕ.+ j ≡ Fin.toℕ off ℕ.+ (j ℕ.+ k) - slicedSize i j off = k , (begin - i ℕ.+ j ≡˘⟨ ≡.cong (ℕ._+ j) (P.proj₂ off+k≤i) ⟩ - (Fin.toℕ off ℕ.+ k) ℕ.+ j ≡⟨ ℕₚ.+-assoc (Fin.toℕ off) k j ⟩ - Fin.toℕ off ℕ.+ (k ℕ.+ j) ≡⟨ ≡.cong (Fin.toℕ off ℕ.+_) (ℕₚ.+-comm k j) ⟩ - Fin.toℕ off ℕ.+ (j ℕ.+ k) ∎) + slicedSize : ∀ n m (i : Fin (suc n)) → P.∃ λ k → n ℕ.+ m ≡ Fin.toℕ i ℕ.+ (m ℕ.+ k) × Fin.toℕ i ℕ.+ k ≡ n + slicedSize n m i = k , (begin + n ℕ.+ m ≡˘⟨ ≡.cong (ℕ._+ m) (P.proj₂ i+k≡n) ⟩ + (Fin.toℕ i ℕ.+ k) ℕ.+ m ≡⟨ ℕₚ.+-assoc (Fin.toℕ i) k m ⟩ + Fin.toℕ i ℕ.+ (k ℕ.+ m) ≡⟨ ≡.cong (Fin.toℕ i ℕ.+_) (ℕₚ.+-comm k m) ⟩ + Fin.toℕ i ℕ.+ (m ℕ.+ k) ∎) , + P.proj₂ i+k≡n where open ≡-Reasoning - off+k≤i = m≤n⇒m+k≡n (ℕₚ.≤-pred (Finₚ.toℕ