Documentation

Init.Data.Fin.Iterate

def Fin.hIterateFrom (P : NatSort u_1) {n : Nat} (f : (i : Fin n) → P iP (i + 1)) (i : Nat) (ubnd : i n) (a : P i) :
P n

hIterateFrom f i bnd a applies f over indices [i:n] to compute P n from P i.

See hIterate below for more details.

Equations
Instances For
    def Fin.hIterate (P : NatSort u_1) {n : Nat} (init : P 0) (f : (i : Fin n) → P iP (i + 1)) :
    P n

    hIterate is a heterogenous iterative operation that applies a index-dependent function f to a value init : P start a total of stop - start times to produce a value of type P stop.

    Concretely, hIterate start stop f init is equal to

      init |> f start _ |> f (start+1) _ ... |> f (end-1) _
    

    Because it is heterogenous and must return a value of type P stop, hIterate requires proof that start ≤ stop.

    One can prove properties of hIterate using the general theorem hIterate_elim or other more specialized theorems.

    Equations
    Instances For
      theorem Fin.hIterate_elim {P : NatSort u_1} (Q : (i : Nat) → P iProp) {n : Nat} (f : (i : Fin n) → P iP (i + 1)) (s : P 0) (init : Q 0 s) (step : ∀ (k : Fin n) (s : P k), Q (k) sQ (k + 1) (f k s)) :
      Q n (Fin.hIterate P s f)
      theorem Fin.hIterate_eq {P : NatSort u_1} (state : (i : Nat) → P i) {n : Nat} (f : (i : Fin n) → P iP (i + 1)) (s : P 0) (init : s = state 0) (step : ∀ (i : Fin n), f i (state i) = state (i + 1)) :
      Fin.hIterate P s f = state n