Documentation

LeanCamCombi.SliceRank

inductive HasSliceRankLE {ι : Type u_1} {R : Type u_2} {α : ιType u_3} [Semiring R] :
(((i : ι) → α i)R)Prop

A function f has slice-rank at most n if it can be written as the sum of n functions of the form x ↦ g (x i) * h (x 1, ..., x (i - 1), x (i + 1), ..., x k).

Instances For
    theorem hasSliceRankLE_iff {ι : Type u_1} {R : Type u_2} {α : ιType u_3} [Semiring R] (a✝ : ) (a✝¹ : ((i : ι) → α i)R) :
    HasSliceRankLE a✝ a✝¹ a✝ = 0 a✝¹ = 0 ∃ (n : ) (f : ((i : ι) → α i)R) (i : ι) (g : α iR) (h : ((j : ι) → j iα j)R), HasSliceRankLE n f a✝ = n + 1 a✝¹ = f + fun (x : (i : ι) → α i) => g (x i) * h fun (j : ι) (x_1 : j i) => x j
    @[simp]
    theorem hasSliceRankLE_zero {ι : Type u_1} {R : Type u_2} {α : ιType u_3} [Semiring R] {f : ((i : ι) → α i)R} :
    theorem hasSliceRankLE_succ {ι : Type u_1} {R : Type u_2} {α : ιType u_3} [Semiring R] {n : } {f : ((i : ι) → α i)R} :
    HasSliceRankLE (n + 1) f ∃ (f' : ((i : ι) → α i)R) (i : ι) (g : α iR) (h : ((j : ι) → j iα j)R), HasSliceRankLE n f' f = f' + fun (x : (i : ι) → α i) => g (x i) * h fun (j : ι) (x_1 : j i) => x j
    theorem hasSliceRankLE_one {ι : Type u_1} {R : Type u_2} {α : ιType u_3} [Semiring R] {f : ((i : ι) → α i)R} :
    HasSliceRankLE 1 f ∃ (i : ι) (g : α iR) (h : ((j : ι) → j iα j)R), f = fun (x : (i : ι) → α i) => g (x i) * h fun (j : ι) (x_1 : j i) => x j
    theorem hasSliceRankLE_iff_exists_sum {ι : Type u_1} {R : Type u_2} {α : ιType u_3} [Semiring R] {n : } {f : ((i : ι) → α i)R} :
    HasSliceRankLE n f ∃ (i : Fin nι) (g : (k : Fin n) → α (i k)R) (h : (k : Fin n) → ((j : ι) → j i kα j)R), f = k : Fin n, fun (x : (i : ι) → α i) => g k (x (i k)) * h k fun (j : ι) (x_1 : j i k) => x j
    theorem HasSliceRankLE.add {ι : Type u_1} {R : Type u_2} {α : ιType u_3} [Semiring R] {m : } {f₁ : ((i : ι) → α i)R} (h₁ : HasSliceRankLE m f₁) {n : } {f₂ : ((i : ι) → α i)R} :
    HasSliceRankLE n f₂HasSliceRankLE (m + n) (f₁ + f₂)
    theorem hasSliceRankLE_card {ι : Type u_1} {R : Type u_2} {α : ιType u_3} [Semiring R] [DecidableEq ι] [Fintype ι] [(i : ι) → Fintype (α i)] (f : ((i : ι) → α i)R) :
    HasSliceRankLE (Fintype.card ((i : ι) → α i)) f

    Any function has slice-rank bounded by the cardinality of its domain.

    theorem exists_hasSliceRankLE {ι : Type u_1} {R : Type u_2} {α : ιType u_3} [Semiring R] [DecidableEq ι] [Finite ι] [∀ (i : ι), Finite (α i)] (f : ((i : ι) → α i)R) :
    ∃ (n : ), HasSliceRankLE n f

    Any function from a finite type has finite slice-rank.