Documentation

Init.Data.Fin.Fold

@[inline]
def Fin.foldl {α : Sort u_1} (n : Nat) (f : αFin nα) (init : α) :
α

Folds over Fin n from the left: foldl 3 f x = f (f (f x 0) 1) 2.

Instances For
    def Fin.foldl.loop {α : Sort u_1} (n : Nat) (f : αFin nα) (x : α) (i : Nat) :
    α

    Inner loop for Fin.foldl. Fin.foldl.loop n f x i = f (f (f x i) ...) (n-1)

    Equations
    Instances For
      @[inline]
      def Fin.foldr {α : Sort u_1} (n : Nat) (f : Fin nαα) (init : α) :
      α

      Folds over Fin n from the right: foldr 3 f x = f 0 (f 1 (f 2 x)).

      Instances For
        def Fin.foldr.loop {α : Sort u_1} (n : Nat) (f : Fin nαα) (i : Nat) :
        i nαα

        Inner loop for Fin.foldr. Fin.foldr.loop n f i x = f 0 (f ... (f (i-1) x))

        Equations
        Instances For
          @[inline]
          def Fin.foldlM {m : Type u_1 → Type u_2} {α : Type u_1} [Monad m] (n : Nat) (f : αFin nm α) (init : α) :
          m α

          Folds a monadic function over Fin n from left to right:

          Fin.foldlM n f x₀ = do
            let x₁ ← f x₀ 0
            let x₂ ← f x₁ 1
            ...
            let xₙ ← f xₙ₋₁ (n-1)
            pure xₙ
          
          Instances For
            @[irreducible]
            def Fin.foldlM.loop {m : Type u_1 → Type u_2} {α : Type u_1} [Monad m] (n : Nat) (f : αFin nm α) (x : α) (i : Nat) :
            m α

            Inner loop for Fin.foldlM.

            Fin.foldlM.loop n f xᵢ i = do
              let xᵢ₊₁ ← f xᵢ i
              ...
              let xₙ ← f xₙ₋₁ (n-1)
              pure xₙ
            
            Equations
            Instances For
              @[inline]
              def Fin.foldrM {m : Type u_1 → Type u_2} {α : Type u_1} [Monad m] (n : Nat) (f : Fin nαm α) (init : α) :
              m α

              Folds a monadic function over Fin n from right to left:

              Fin.foldrM n f xₙ = do
                let xₙ₋₁ ← f (n-1) xₙ
                let xₙ₋₂ ← f (n-2) xₙ₋₁
                ...
                let x₀ ← f 0 x₁
                pure x₀
              
              Equations
              Instances For
                @[irreducible]
                def Fin.foldrM.loop {m : Type u_1 → Type u_2} {α : Type u_1} [Monad m] (n : Nat) (f : Fin nαm α) :
                { i : Nat // i n }αm α

                Inner loop for Fin.foldrM.

                Fin.foldrM.loop n f i xᵢ = do
                  let xᵢ₋₁ ← f (i-1) xᵢ
                  ...
                  let x₁ ← f 1 x₂
                  let x₀ ← f 0 x₁
                  pure x₀
                
                Equations
                Instances For

                  foldlM #

                  theorem Fin.foldlM_loop_lt {m : Type u_1 → Type u_2} {α : Type u_1} {n i : Nat} [Monad m] (f : αFin nm α) (x : α) (h : i < n) :
                  Fin.foldlM.loop n f x i = do let xf x i, h Fin.foldlM.loop n f x (i + 1)
                  theorem Fin.foldlM_loop_eq {m : Type u_1 → Type u_2} {α : Type u_1} {n : Nat} [Monad m] (f : αFin nm α) (x : α) :
                  @[irreducible]
                  theorem Fin.foldlM_loop {m : Type u_1 → Type u_2} {α : Type u_1} {n i : Nat} [Monad m] (f : αFin (n + 1)m α) (x : α) (h : i < n + 1) :
                  Fin.foldlM.loop (n + 1) f x i = do let xf x i, h Fin.foldlM.loop n (fun (x : α) (j : Fin n) => f x j.succ) x i
                  @[simp]
                  theorem Fin.foldlM_zero {m : Type u_1 → Type u_2} {α : Type u_1} [Monad m] (f : αFin 0m α) (x : α) :
                  Fin.foldlM 0 f x = pure x
                  theorem Fin.foldlM_succ {m : Type u_1 → Type u_2} {α : Type u_1} {n : Nat} [Monad m] (f : αFin (n + 1)m α) (x : α) :
                  Fin.foldlM (n + 1) f x = f x 0 >>= Fin.foldlM n fun (x : α) (j : Fin n) => f x j.succ

                  foldrM #

                  theorem Fin.foldrM_loop_zero {m : Type u_1 → Type u_2} {n : Nat} {α : Type u_1} [Monad m] (f : Fin nαm α) (x : α) :
                  Fin.foldrM.loop n f 0, x = pure x
                  theorem Fin.foldrM_loop_succ {m : Type u_1 → Type u_2} {n : Nat} {α : Type u_1} {i : Nat} [Monad m] (f : Fin nαm α) (x : α) (h : i < n) :
                  Fin.foldrM.loop n f i + 1, h x = f i, h x >>= Fin.foldrM.loop n f i,
                  theorem Fin.foldrM_loop {m : Type u_1 → Type u_2} {n : Nat} {α : Type u_1} {i : Nat} [Monad m] [LawfulMonad m] (f : Fin (n + 1)αm α) (x : α) (h : i + 1 n + 1) :
                  Fin.foldrM.loop (n + 1) f i + 1, h x = Fin.foldrM.loop n (fun (j : Fin n) => f j.succ) i, x >>= f 0
                  @[simp]
                  theorem Fin.foldrM_zero {m : Type u_1 → Type u_2} {α : Type u_1} [Monad m] (f : Fin 0αm α) (x : α) :
                  Fin.foldrM 0 f x = pure x
                  theorem Fin.foldrM_succ {m : Type u_1 → Type u_2} {n : Nat} {α : Type u_1} [Monad m] [LawfulMonad m] (f : Fin (n + 1)αm α) (x : α) :
                  Fin.foldrM (n + 1) f x = Fin.foldrM n (fun (i : Fin n) => f i.succ) x >>= f 0

                  foldl #

                  theorem Fin.foldl_loop_lt {α : Sort u_1} {n i : Nat} (f : αFin nα) (x : α) (h : i < n) :
                  Fin.foldl.loop n f x i = Fin.foldl.loop n f (f x i, h) (i + 1)
                  theorem Fin.foldl_loop_eq {α : Sort u_1} {n : Nat} (f : αFin nα) (x : α) :
                  Fin.foldl.loop n f x n = x
                  @[irreducible]
                  theorem Fin.foldl_loop {α : Sort u_1} {n i : Nat} (f : αFin (n + 1)α) (x : α) (h : i < n + 1) :
                  Fin.foldl.loop (n + 1) f x i = Fin.foldl.loop n (fun (x : α) (j : Fin n) => f x j.succ) (f x i, h) i
                  @[simp]
                  theorem Fin.foldl_zero {α : Sort u_1} (f : αFin 0α) (x : α) :
                  Fin.foldl 0 f x = x
                  theorem Fin.foldl_succ {α : Sort u_1} {n : Nat} (f : αFin (n + 1)α) (x : α) :
                  Fin.foldl (n + 1) f x = Fin.foldl n (fun (x : α) (i : Fin n) => f x i.succ) (f x 0)
                  theorem Fin.foldl_succ_last {α : Sort u_1} {n : Nat} (f : αFin (n + 1)α) (x : α) :
                  Fin.foldl (n + 1) f x = f (Fin.foldl n (fun (x1 : α) (x2 : Fin n) => f x1 x2.castSucc) x) (Fin.last n)
                  theorem Fin.foldl_eq_foldlM {α : Type u_1} {n : Nat} (f : αFin nα) (x : α) :
                  Fin.foldl n f x = Fin.foldlM n f x

                  foldr #

                  theorem Fin.foldr_loop_zero {n : Nat} {α : Sort u_1} (f : Fin nαα) (x : α) :
                  Fin.foldr.loop n f 0 x = x
                  theorem Fin.foldr_loop_succ {n : Nat} {α : Sort u_1} {i : Nat} (f : Fin nαα) (x : α) (h : i < n) :
                  Fin.foldr.loop n f (i + 1) h x = Fin.foldr.loop n f i (f i, h x)
                  theorem Fin.foldr_loop {n : Nat} {α : Sort u_1} {i : Nat} (f : Fin (n + 1)αα) (x : α) (h : i + 1 n + 1) :
                  Fin.foldr.loop (n + 1) f (i + 1) h x = f 0 (Fin.foldr.loop n (fun (j : Fin n) => f j.succ) i x)
                  @[simp]
                  theorem Fin.foldr_zero {α : Sort u_1} (f : Fin 0αα) (x : α) :
                  Fin.foldr 0 f x = x
                  theorem Fin.foldr_succ {n : Nat} {α : Sort u_1} (f : Fin (n + 1)αα) (x : α) :
                  Fin.foldr (n + 1) f x = f 0 (Fin.foldr n (fun (i : Fin n) => f i.succ) x)
                  theorem Fin.foldr_succ_last {n : Nat} {α : Sort u_1} (f : Fin (n + 1)αα) (x : α) :
                  Fin.foldr (n + 1) f x = Fin.foldr n (fun (x : Fin n) => f x.castSucc) (f (Fin.last n) x)
                  theorem Fin.foldr_eq_foldrM {n : Nat} {α : Type u_1} (f : Fin nαα) (x : α) :
                  Fin.foldr n f x = Fin.foldrM n f x
                  theorem Fin.foldl_rev {n : Nat} {α : Sort u_1} (f : Fin nαα) (x : α) :
                  Fin.foldl n (fun (x : α) (i : Fin n) => f i.rev x) x = Fin.foldr n f x
                  theorem Fin.foldr_rev {α : Sort u_1} {n : Nat} (f : αFin nα) (x : α) :
                  Fin.foldr n (fun (i : Fin n) (x : α) => f x i.rev) x = Fin.foldl n f x