Documentation

Mathlib.Data.Fin.Basic

The finite type with n elements #

Fin n is the type whose elements are natural numbers smaller than n. This file expands on the development in the core library.

Main definitions #

Induction principles #

Order embeddings and an order isomorphism #

Other casts #

Misc definitions #

def finZeroElim {α : Fin 0Sort u_1} (x : Fin 0) :
α x

Elimination principle for the empty set Fin 0, dependent version.

Equations
Instances For
    def Fin.elim0' {α : Sort u_1} (x : Fin 0) :
    α

    A non-dependent variant of elim0.

    Equations
    Instances For
      theorem Fin.size_positive {n : } :
      Fin n0 < n

      If you actually have an element of Fin n, then the n is always positive

      theorem Fin.size_positive' {n : } [Nonempty (Fin n)] :
      0 < n
      theorem Fin.prop {n : } (a : Fin n) :
      a < n
      @[simp]
      theorem Fin.equivSubtype_apply {n : } (a : Fin n) :
      Fin.equivSubtype a = { val := a, property := (_ : a < n) }
      @[simp]
      theorem Fin.equivSubtype_symm_apply {n : } (a : { i // i < n }) :
      Fin.equivSubtype.symm a = { val := a, isLt := (_ : a < n) }
      def Fin.equivSubtype {n : } :
      Fin n { i // i < n }

      Equivalence between Fin n and { i // i < n }.

      Equations
      • One or more equations did not get rendered due to their size.
      Instances For

        coercions and constructions #

        theorem Fin.val_eq_val {n : } (a : Fin n) (b : Fin n) :
        a = b a = b
        theorem Fin.eq_iff_veq {n : } (a : Fin n) (b : Fin n) :
        a = b a = b
        theorem Fin.ne_iff_vne {n : } (a : Fin n) (b : Fin n) :
        a b a b
        @[simp]
        theorem Fin.mk_eq_mk {n : } {a : } {h : a < n} {a' : } {h' : a' < n} :
        { val := a, isLt := h } = { val := a', isLt := h' } a = a'
        theorem Fin.heq_fun_iff {α : Sort u_1} {k : } {l : } (h : k = l) {f : Fin kα} {g : Fin lα} :
        HEq f g ∀ (i : Fin k), f i = g { val := i, isLt := (_ : i < l) }

        Assume k = l. If two functions defined on Fin k and Fin l are equal on each element, then they coincide (in the heq sense).

        theorem Fin.heq_fun₂_iff {α : Sort u_1} {k : } {l : } {k' : } {l' : } (h : k = l) (h' : k' = l') {f : Fin kFin k'α} {g : Fin lFin l'α} :
        HEq f g ∀ (i : Fin k) (j : Fin k'), f i j = g { val := i, isLt := (_ : i < l) } { val := j, isLt := (_ : j < l') }

        Assume k = l and k' = l'. If two functions Fin k → Fin k' → α and Fin l → Fin l' → α are equal on each pair, then they coincide (in the heq sense).

        theorem Fin.heq_ext_iff {k : } {l : } (h : k = l) {i : Fin k} {j : Fin l} :
        HEq i j i = j

        order #

        theorem Fin.lt_iff_val_lt_val {n : } {a : Fin n} {b : Fin n} :
        a < b a < b
        theorem Fin.le_iff_val_le_val {n : } {a : Fin n} {b : Fin n} :
        a b a b
        @[simp]
        theorem Fin.val_fin_lt {n : } {a : Fin n} {b : Fin n} :
        a < b a < b

        a < b as natural numbers if and only if a < b in Fin n.

        @[simp]
        theorem Fin.val_fin_le {n : } {a : Fin n} {b : Fin n} :
        a b a b

        a ≤ b as natural numbers if and only if a ≤ b in Fin n.

        Equations
        • One or more equations did not get rendered due to their size.
        theorem Fin.min_val {n : } {a : Fin n} :
        min (a) n = a
        theorem Fin.max_val {n : } {a : Fin n} :
        max (a) n = n
        Equations
        • Fin.instPartialOrderFin = inferInstance
        theorem Fin.val_strictMono {n : } :
        StrictMono Fin.val
        @[simp]
        theorem Fin.orderIsoSubtype_apply {n : } (a : Fin n) :
        Fin.orderIsoSubtype a = { val := a, property := (_ : a < n) }
        @[simp]
        theorem Fin.orderIsoSubtype_symm_apply {n : } (a : { i // i < n }) :
        ↑(RelIso.symm Fin.orderIsoSubtype) a = { val := a, isLt := (_ : a < n) }
        def Fin.orderIsoSubtype {n : } :
        Fin n ≃o { i // i < n }

        The equivalence Fin n ≃ { i // i < n } is an order isomorphism.

        Equations
        • One or more equations did not get rendered due to their size.
        Instances For
          @[simp]
          theorem Fin.valEmbedding_apply {n : } (self : Fin n) :
          Fin.valEmbedding self = self

          The inclusion map Fin n → ℕ is an embedding.

          Equations
          Instances For
            @[simp]
            theorem Fin.equivSubtype_symm_trans_valEmbedding {n : } :
            Function.Embedding.trans (Equiv.toEmbedding Fin.equivSubtype.symm) Fin.valEmbedding = Function.Embedding.subtype fun x => x < n
            @[simp]
            theorem Fin.valOrderEmbedding_apply (n : ) (self : Fin n) :
            ↑(Fin.valOrderEmbedding n) self = self

            The inclusion map Fin n → ℕ is an order embedding.

            Equations
            • Fin.valOrderEmbedding n = { toEmbedding := Fin.valEmbedding, map_rel_iff' := (_ : ∀ {a b : Fin n}, Fin.valEmbedding a Fin.valEmbedding b Fin.valEmbedding a Fin.valEmbedding b) }
            Instances For
              instance Fin.Lt.isWellOrder (n : ) :
              IsWellOrder (Fin n) fun x x_1 => x < x_1

              The ordering on Fin n is a well order.

              Equations

              Use the ordering on Fin n for checking recursive definitions.

              For example, the following definition is not accepted by the termination checker, unless we declare the WellFoundedRelation instance:

              def factorial {n : ℕ} : Fin n → ℕ
                | ⟨0, _⟩ := 1
                | ⟨i + 1, hi⟩ := (i + 1) * factorial ⟨i, i.lt_succ_self.trans hi⟩
              
              Equations
              • Fin.instWellFoundedRelationFin = measure Fin.val
              def Fin.ofNat'' {n : } [NeZero n] (i : ) :
              Fin n

              Given a positive n, Fin.ofNat' i is i % n as an element of Fin n.

              Equations
              Instances For
                instance Fin.instZeroFin {n : } [NeZero n] :
                Zero (Fin n)
                Equations
                instance Fin.instOneFin {n : } [NeZero n] :
                One (Fin n)
                Equations
                @[simp]
                theorem Fin.val_zero' (n : ) [NeZero n] :
                0 = 0

                The Fin.val_zero in Std only applies in Fin (n+1). This one instead uses a NeZero n typeclass hypothesis.

                @[simp]
                theorem Fin.zero_le' {n : } [NeZero n] (a : Fin n) :
                0 a

                The Fin.zero_le in Std only applies in Fin (n+1). This one instead uses a NeZero n typeclass hypothesis.

                theorem Fin.pos_iff_ne_zero' {n : } [NeZero n] (a : Fin n) :
                0 < a a 0

                The Fin.pos_iff_ne_zero in Std only applies in Fin (n+1). This one instead uses a NeZero n typeclass hypothesis.

                @[simp]
                theorem Fin.revPerm_apply {n : } (i : Fin n) :
                Fin.revPerm i = Fin.rev i
                @[simp]
                theorem Fin.revPerm_symm_apply {n : } (i : Fin n) :
                Fin.revPerm.symm i = Fin.rev i
                def Fin.revPerm {n : } :

                Fin.rev as an Equiv.Perm, the antitone involution Fin n → Fin n given by i ↦ n-(i+1).

                Equations
                Instances For
                  @[simp]
                  theorem Fin.revPerm_symm {n : } :
                  Fin.revPerm.symm = Fin.revPerm
                  @[simp]
                  theorem Fin.revOrderIso_toEquiv {n : } :
                  Fin.revOrderIso.toEquiv = OrderDual.ofDual.trans Fin.revPerm
                  @[simp]
                  theorem Fin.revOrderIso_apply {n : } :
                  ∀ (a : (Fin n)ᵒᵈ), Fin.revOrderIso a = Fin.rev (OrderDual.ofDual a)

                  Fin.rev n as an order-reversing isomorphism.

                  Equations
                  • One or more equations did not get rendered due to their size.
                  Instances For
                    @[simp]
                    theorem Fin.revOrderIso_symm_apply {n : } (i : Fin n) :
                    ↑(OrderIso.symm Fin.revOrderIso) i = OrderDual.toDual (Fin.rev i)
                    Equations
                    • Fin.instBoundedOrderFinHAddNatInstHAddInstAddNatOfNatInstLEFin = BoundedOrder.mk
                    Equations
                    • Fin.instLatticeFinHAddNatInstHAddInstAddNatOfNat = LinearOrder.toLattice
                    theorem Fin.bot_eq_zero (n : ) :
                    = 0
                    @[simp]
                    theorem Fin.coe_orderIso_apply {n : } {m : } (e : Fin n ≃o Fin m) (i : Fin n) :
                    ↑(e i) = i

                    If e is an orderIso between Fin n and Fin m, then n = m and e is the identity map. In this lemma we state that for each i : Fin n we have (e i : ℕ) = (i : ℕ).

                    instance Fin.orderIsoUnique {n : } :
                    Equations
                    theorem Fin.strictMono_unique {n : } {α : Type u_1} [Preorder α] {f : Fin nα} {g : Fin nα} (hf : StrictMono f) (hg : StrictMono g) (h : Set.range f = Set.range g) :
                    f = g

                    Two strictly monotone functions from Fin n are equal provided that their ranges are equal.

                    theorem Fin.orderEmbedding_eq {n : } {α : Type u_1} [Preorder α] {f : Fin n ↪o α} {g : Fin n ↪o α} (h : Set.range f = Set.range g) :
                    f = g

                    Two order embeddings of Fin n are equal provided that their ranges are equal.

                    addition, numerals, and coercion from Nat #

                    @[simp]
                    theorem Fin.val_one' (n : ) [NeZero n] :
                    1 = 1 % n
                    theorem Fin.val_one'' {n : } :
                    1 = 1 % (n + 1)
                    Equations
                    • One or more equations did not get rendered due to their size.
                    theorem Fin.add_zero {n : } [NeZero n] (k : Fin n) :
                    k + 0 = k
                    theorem Fin.zero_add {n : } [NeZero n] (k : Fin n) :
                    0 + k = k
                    instance Fin.instOfNatFin {n : } {a : } [NeZero n] :
                    OfNat (Fin n) a
                    Equations
                    @[simp]
                    theorem Fin.ofNat'_zero {n : } {h : n > 0} [NeZero n] :
                    @[simp]
                    theorem Fin.ofNat'_one {n : } {h : n > 0} [NeZero n] :
                    Equations
                    instance Fin.addCommMonoid (n : ) [NeZero n] :
                    Equations
                    Equations
                    theorem Fin.val_add_eq_ite {n : } (a : Fin n) (b : Fin n) :
                    ↑(a + b) = if n a + b then a + b - n else a + b
                    @[deprecated]
                    theorem Fin.val_bit0 {n : } (k : Fin n) :
                    ↑(bit0 k) = bit0 k % n
                    @[deprecated]
                    theorem Fin.val_bit1 {n : } [NeZero n] (k : Fin n) :
                    ↑(bit1 k) = bit1 k % n
                    @[simp, deprecated]
                    theorem Fin.mk_bit0 {m : } {n : } (h : bit0 m < n) :
                    { val := bit0 m, isLt := h } = bit0 { val := m, isLt := (_ : m < n) }
                    @[simp, deprecated]
                    theorem Fin.mk_bit1 {m : } {n : } [NeZero n] (h : bit1 m < n) :
                    { val := bit1 m, isLt := h } = bit1 { val := m, isLt := (_ : m < n) }
                    @[simp]
                    theorem Fin.ofNat_eq_val (n : ) [NeZero n] (a : ) :
                    Fin.ofNat'' a = a
                    theorem Fin.val_cast_of_lt {n : } [NeZero n] {a : } (h : a < n) :
                    a = a

                    Converting an in-range number to Fin (n + 1) produces a result whose value is the original number.

                    theorem Fin.cast_val_eq_self {n : } [NeZero n] (a : Fin n) :
                    a = a

                    Converting the value of a Fin (n + 1) to Fin (n + 1) results in the same value.

                    theorem Fin.le_val_last {n : } (i : Fin (n + 1)) :
                    i n
                    @[simp]
                    theorem Fin.zero_eq_one_iff {n : } [NeZero n] :
                    0 = 1 n = 1
                    @[simp]
                    theorem Fin.one_eq_zero_iff {n : } [NeZero n] :
                    1 = 0 n = 1

                    succ and casts into larger Fin types #

                    def Fin.succEmbedding (n : ) :
                    Fin n ↪o Fin (n + 1)

                    Fin.succ as an OrderEmbedding

                    Equations
                    Instances For
                      @[simp]
                      theorem Fin.val_succEmbedding {n : } :
                      ↑(Fin.succEmbedding n) = Fin.succ
                      @[simp]
                      theorem Fin.succ_zero_eq_one' {n : } [NeZero n] :
                      @[simp]
                      theorem Fin.succ_one_eq_two' {n : } [NeZero n] :

                      The Fin.succ_one_eq_two in Std only applies in Fin (n+2). This one instead uses a NeZero n typeclass hypothesis.

                      @[simp]
                      theorem Fin.le_zero_iff' {n : } [NeZero n] {k : Fin n} :
                      k 0 k = 0

                      The Fin.le_zero_iff in Std only applies in Fin (n+1). This one instead uses a NeZero n typeclass hypothesis.

                      @[simp]
                      theorem Fin.cast_refl {n : } (h : n = n) :
                      Fin.cast h = id
                      theorem Fin.strictMono_castLE {n : } {m : } (h : n m) :
                      @[simp]
                      theorem Fin.castLEEmb_toEmbedding {n : } {m : } (h : n m) :
                      (Fin.castLEEmb h).toEmbedding = { toFun := Fin.castLE h, inj' := (_ : ∀ (x x_1 : Fin n), Fin.castLE h x = Fin.castLE h x_1x = x_1) }
                      @[simp]
                      theorem Fin.castLEEmb_apply {n : } {m : } (h : n m) (i : Fin n) :
                      def Fin.castLEEmb {n : } {m : } (h : n m) :

                      Fin.castLE as an OrderEmbedding, castLEEmb h i embeds i into a larger Fin type.

                      Equations
                      Instances For
                        @[simp]
                        theorem Fin.range_castLE {n : } {k : } (h : n k) :
                        Set.range (Fin.castLE h) = {i | i < n}
                        @[simp]
                        theorem Fin.coe_of_injective_castLEEmb_symm {n : } {k : } (h : n k) (i : Fin k) (hi : i Set.range ↑(Fin.castLEEmb h)) :
                        ↑((Equiv.ofInjective ↑(Fin.castLEEmb h) (_ : Function.Injective ↑(Fin.castLEEmb h))).symm { val := i, property := hi }) = i
                        theorem Fin.leftInverse_cast {n : } {m : } (eq : n = m) :
                        theorem Fin.rightInverse_cast {n : } {m : } (eq : n = m) :
                        theorem Fin.cast_le_cast {n : } {m : } (eq : n = m) {a : Fin n} {b : Fin n} :
                        Fin.cast eq a Fin.cast eq b a b
                        @[simp]
                        theorem Fin.castIso_apply {n : } {m : } (eq : n = m) (i : Fin n) :
                        ↑(Fin.castIso eq) i = Fin.cast eq i
                        @[simp]
                        theorem Fin.castIso_symm_apply {n : } {m : } (eq : n = m) (i : Fin m) :
                        ↑(RelIso.symm (Fin.castIso eq)) i = Fin.cast (_ : m = n) i
                        def Fin.castIso {n : } {m : } (eq : n = m) :

                        Fin.cast as an OrderIso, castIso eq i embeds i into an equal Fin type, see also Equiv.finCongr.

                        Equations
                        • One or more equations did not get rendered due to their size.
                        Instances For
                          @[simp]
                          theorem Fin.symm_castIso {n : } {m : } (h : n = m) :
                          @[simp]
                          theorem Fin.cast_zero {n : } {n' : } [NeZero n] {h : n = n'} :
                          Fin.cast h 0 = 0
                          @[simp]
                          theorem Fin.castIso_refl {n : } (h : optParam (n = n) (_ : n = n)) :
                          theorem Fin.castIso_to_equiv {n : } {m : } (h : n = m) :
                          (Fin.castIso h).toEquiv = Equiv.cast (_ : Fin n = Fin m)

                          While in many cases Fin.castIso is better than Equiv.cast/cast, sometimes we want to apply a generic theorem about cast.

                          theorem Fin.cast_eq_cast {n : } {m : } (h : n = m) :
                          Fin.cast h = cast (_ : Fin n = Fin m)

                          While in many cases Fin.cast is better than Equiv.cast/cast, sometimes we want to apply a generic theorem about cast.

                          @[simp]
                          theorem Fin.castAddEmb_apply {n : } (m : ) :
                          ∀ (a : Fin n), ↑(Fin.castAddEmb m) a = Fin.castAdd m a
                          @[simp]
                          theorem Fin.castAddEmb_toEmbedding {n : } (m : ) :
                          (Fin.castAddEmb m).toEmbedding = { toFun := Fin.castAdd m, inj' := (_ : ∀ (x x_1 : Fin n), Fin.castAdd m x = Fin.castAdd m x_1x = x_1) }
                          def Fin.castAddEmb {n : } (m : ) :
                          Fin n ↪o Fin (n + m)

                          Fin.castAdd as an OrderEmbedding, castAddEmb m i embeds i : Fin n in Fin (n+m). See also Fin.natAddEmb and Fin.addNatEmb.

                          Equations
                          Instances For
                            theorem Fin.strictMono_castSucc {n : } :
                            StrictMono Fin.castSucc
                            @[simp]
                            theorem Fin.castSuccEmb_toEmbedding {n : } :
                            Fin.castSuccEmb.toEmbedding = { toFun := Fin.castSucc, inj' := (_ : ∀ (x x_1 : Fin n), Fin.castSucc x = Fin.castSucc x_1x = x_1) }
                            @[simp]
                            theorem Fin.castSuccEmb_apply {n : } :
                            ∀ (a : Fin n), Fin.castSuccEmb a = Fin.castSucc a
                            def Fin.castSuccEmb {n : } :
                            Fin n ↪o Fin (n + 1)

                            Fin.castSucc as an OrderEmbedding, castSuccEmb i embeds i : Fin n in Fin (n+1).

                            Equations
                            Instances For
                              @[simp]
                              theorem Fin.castSucc_zero' {n : } [NeZero n] :

                              The Fin.castSucc_zero in Std only applies in Fin (n+1). This one instead uses a NeZero n typeclass hypothesis.

                              theorem Fin.castSucc_pos' {n : } [NeZero n] {i : Fin n} (h : 0 < i) :

                              castSucc i is positive when i is positive.

                              The Fin.castSucc_pos in Std only applies in Fin (n+1). This one instead uses a NeZero n typeclass hypothesis.

                              @[simp]
                              theorem Fin.castSucc_eq_zero_iff' {n : } [NeZero n] (a : Fin n) :

                              The Fin.castSucc_eq_zero_iff in Std only applies in Fin (n+1). This one instead uses a NeZero n typeclass hypothesis.

                              theorem Fin.castSucc_ne_zero_iff' {n : } [NeZero n] (a : Fin n) :

                              The Fin.castSucc_ne_zero_iff in Std only applies in Fin (n+1). This one instead uses a NeZero n typeclass hypothesis.

                              @[simp]
                              theorem Fin.coe_eq_castSucc {n : } {a : Fin n} :
                              a = Fin.castSucc a
                              @[simp]
                              theorem Fin.range_castSucc {n : } :
                              Set.range Fin.castSucc = {i | i < n}
                              @[simp]
                              theorem Fin.coe_of_injective_castSucc_symm {n : } (i : Fin (Nat.succ n)) (hi : i Set.range Fin.castSucc) :
                              ↑((Equiv.ofInjective Fin.castSucc (_ : Function.Injective Fin.castSucc)).symm { val := i, property := hi }) = i
                              theorem Fin.strictMono_addNat {n : } (m : ) :
                              StrictMono fun x => Fin.addNat x m
                              @[simp]
                              theorem Fin.addNatEmb_apply {n : } (m : ) :
                              ∀ (x : Fin n), ↑(Fin.addNatEmb m) x = Fin.addNat x m
                              @[simp]
                              theorem Fin.addNatEmb_toEmbedding {n : } (m : ) :
                              (Fin.addNatEmb m).toEmbedding = { toFun := fun x => Fin.addNat x m, inj' := (_ : ∀ (x x_1 : Fin n), Fin.addNat x m = Fin.addNat x_1 mx = x_1) }
                              def Fin.addNatEmb {n : } (m : ) :
                              Fin n ↪o Fin (n + m)

                              Fin.addNat as an OrderEmbedding, addNatEmb m i adds m to i, generalizes Fin.succ.

                              Equations
                              Instances For
                                @[simp]
                                theorem Fin.natAddEmb_apply (n : ) {m : } (i : Fin m) :
                                @[simp]
                                theorem Fin.natAddEmb_toEmbedding (n : ) {m : } :
                                (Fin.natAddEmb n).toEmbedding = { toFun := Fin.natAdd n, inj' := (_ : ∀ (x x_1 : Fin m), Fin.natAdd n x = Fin.natAdd n x_1x = x_1) }
                                def Fin.natAddEmb (n : ) {m : } :
                                Fin m ↪o Fin (n + m)

                                Fin.natAdd as an OrderEmbedding, natAddEmb n i adds n to i "on the left".

                                Equations
                                Instances For

                                  pred #

                                  def Fin.divNat {n : } {m : } (i : Fin (m * n)) :
                                  Fin m

                                  Compute i / n, where n is a Nat and inferred the type of i.

                                  Equations
                                  Instances For
                                    @[simp]
                                    theorem Fin.coe_divNat {n : } {m : } (i : Fin (m * n)) :
                                    ↑(Fin.divNat i) = i / n
                                    def Fin.modNat {n : } {m : } (i : Fin (m * n)) :
                                    Fin n

                                    Compute i % n, where n is a Nat and inferred the type of i.

                                    Equations
                                    Instances For
                                      @[simp]
                                      theorem Fin.coe_modNat {n : } {m : } (i : Fin (m * n)) :
                                      ↑(Fin.modNat i) = i % n

                                      recursion and induction principles #

                                      theorem Fin.liftFun_iff_succ {n : } {α : Type u_1} (r : ααProp) [IsTrans α r] {f : Fin (n + 1)α} :
                                      ((fun x x_1 => x < x_1) r) f f (i : Fin n) → r (f (Fin.castSucc i)) (f (Fin.succ i))
                                      theorem Fin.strictMono_iff_lt_succ {n : } {α : Type u_1} [Preorder α] {f : Fin (n + 1)α} :
                                      StrictMono f ∀ (i : Fin n), f (Fin.castSucc i) < f (Fin.succ i)

                                      A function f on Fin (n + 1) is strictly monotone if and only if f i < f (i + 1) for all i.

                                      theorem Fin.monotone_iff_le_succ {n : } {α : Type u_1} [Preorder α] {f : Fin (n + 1)α} :
                                      Monotone f ∀ (i : Fin n), f (Fin.castSucc i) f (Fin.succ i)

                                      A function f on Fin (n + 1) is monotone if and only if f i ≤ f (i + 1) for all i.

                                      theorem Fin.strictAnti_iff_succ_lt {n : } {α : Type u_1} [Preorder α] {f : Fin (n + 1)α} :
                                      StrictAnti f ∀ (i : Fin n), f (Fin.succ i) < f (Fin.castSucc i)

                                      A function f on Fin (n + 1) is strictly antitone if and only if f (i + 1) < f i for all i.

                                      theorem Fin.antitone_iff_succ_le {n : } {α : Type u_1} [Preorder α] {f : Fin (n + 1)α} :
                                      Antitone f ∀ (i : Fin n), f (Fin.succ i) f (Fin.castSucc i)

                                      A function f on Fin (n + 1) is antitone if and only if f (i + 1) ≤ f i for all i.

                                      instance Fin.neg (n : ) :
                                      Neg (Fin n)

                                      Negation on Fin n

                                      Equations
                                      • Fin.neg n = { neg := fun a => { val := (n - a) % n, isLt := (_ : (n - a) % n < n) } }
                                      instance Fin.addCommGroup (n : ) [NeZero n] :

                                      Abelian group structure on Fin n.

                                      Equations

                                      Note this is more general than Fin.addCommGroup as it applies (vacuously) to Fin 0 too.

                                      Equations

                                      Note this is more general than Fin.addCommGroup as it applies (vacuously) to Fin 0 too.

                                      Equations

                                      Note this is more general than Fin.addCommGroup as it applies (vacuously) to Fin 0 too.

                                      Equations

                                      Note this is more general than Fin.addCommGroup as it applies (vacuously) to Fin 0 too.

                                      Equations
                                      theorem Fin.coe_neg {n : } (a : Fin n) :
                                      ↑(-a) = (n - a) % n
                                      theorem Fin.coe_sub {n : } (a : Fin n) (b : Fin n) :
                                      ↑(a - b) = (a + (n - b)) % n
                                      @[simp]
                                      theorem Fin.coe_fin_one (a : Fin 1) :
                                      a = 0
                                      @[simp]
                                      theorem Fin.coe_neg_one {n : } :
                                      ↑(-1) = n
                                      theorem Fin.coe_sub_one {n : } (a : Fin (n + 1)) :
                                      ↑(a - 1) = if a = 0 then n else a - 1
                                      theorem Fin.coe_sub_iff_le {n : } {a : Fin n} {b : Fin n} :
                                      ↑(a - b) = a - b b a
                                      theorem Fin.coe_sub_iff_lt {n : } {a : Fin n} {b : Fin n} :
                                      ↑(a - b) = n + a - b a < b
                                      @[simp]
                                      theorem Fin.lt_sub_one_iff {n : } {k : Fin (n + 2)} :
                                      k < k - 1 k = 0
                                      @[simp]
                                      theorem Fin.le_sub_one_iff {n : } {k : Fin (n + 1)} :
                                      k k - 1 k = 0
                                      @[simp]
                                      theorem Fin.sub_one_lt_iff {n : } {k : Fin (n + 1)} :
                                      k - 1 < k 0 < k
                                      theorem Fin.last_sub {n : } (i : Fin (n + 1)) :
                                      def Fin.succAbove {n : } (p : Fin (n + 1)) (i : Fin n) :
                                      Fin (n + 1)

                                      succAbove p i embeds Fin n into Fin (n + 1) with a hole around p.

                                      Equations
                                      Instances For
                                        @[simp]
                                        theorem Fin.succAboveEmb_apply {n : } (p : Fin (n + 1)) (i : Fin n) :
                                        @[simp]
                                        theorem Fin.succAboveEmb_toEmbedding {n : } (p : Fin (n + 1)) :
                                        (Fin.succAboveEmb p).toEmbedding = { toFun := Fin.succAbove p, inj' := (_ : ∀ (x x_1 : Fin n), Fin.succAbove p x = Fin.succAbove p x_1x = x_1) }
                                        def Fin.succAboveEmb {n : } (p : Fin (n + 1)) :
                                        Fin n ↪o Fin (n + 1)

                                        Fin.auccAbove as an OrderEmbedding, succAboveEmb p i embeds Fin n into Fin (n + 1) with a hole around p.

                                        Equations
                                        Instances For
                                          theorem Fin.succAbove_below {n : } (p : Fin (n + 1)) (i : Fin n) (h : Fin.castSucc i < p) :

                                          Embedding i : Fin n into Fin (n + 1) with a hole around p : Fin (n + 1) embeds i by castSucc when the resulting i.castSucc < p.

                                          @[simp]
                                          theorem Fin.succAbove_ne_zero_zero {n : } [NeZero n] {a : Fin (n + 1)} (ha : a 0) :
                                          theorem Fin.succAbove_eq_zero_iff {n : } [NeZero n] {a : Fin (n + 1)} {b : Fin n} (ha : a 0) :
                                          Fin.succAbove a b = 0 b = 0
                                          theorem Fin.succAbove_ne_zero {n : } [NeZero n] {a : Fin (n + 1)} {b : Fin n} (ha : a 0) (hb : b 0) :
                                          @[simp]
                                          theorem Fin.succAbove_zero {n : } :
                                          Fin.succAbove 0 = Fin.succ

                                          Embedding Fin n into Fin (n + 1) with a hole around zero embeds by succ.

                                          @[simp]
                                          theorem Fin.succAbove_last {n : } :
                                          Fin.succAbove (Fin.last n) = Fin.castSucc

                                          Embedding Fin n into Fin (n + 1) with a hole around last n embeds by castSucc.

                                          theorem Fin.succAbove_above {n : } (p : Fin (n + 1)) (i : Fin n) (h : p Fin.castSucc i) :

                                          Embedding i : Fin n into Fin (n + 1) with a hole around p : Fin (n + 1) embeds i by succ when the resulting p < i.succ.

                                          theorem Fin.succAbove_lt_ge {n : } (p : Fin (n + 1)) (i : Fin n) :

                                          Embedding i : Fin n into Fin (n + 1) is always about some hole p.

                                          theorem Fin.succAbove_lt_gt {n : } (p : Fin (n + 1)) (i : Fin n) :

                                          Embedding i : Fin n into Fin (n + 1) is always about some hole p.

                                          @[simp]
                                          theorem Fin.succAbove_lt_iff {n : } (p : Fin (n + 1)) (i : Fin n) :

                                          Embedding i : Fin n into Fin (n + 1) using a pivot p that is greater results in a value that is less than p.

                                          theorem Fin.lt_succAbove_iff {n : } (p : Fin (n + 1)) (i : Fin n) :

                                          Embedding i : Fin n into Fin (n + 1) using a pivot p that is lesser results in a value that is greater than p.

                                          theorem Fin.succAbove_ne {n : } (p : Fin (n + 1)) (i : Fin n) :

                                          Embedding i : Fin n into Fin (n + 1) with a hole around p : Fin (n + 1) never results in p itself

                                          theorem Fin.succAbove_pos {n : } [NeZero n] (p : Fin (n + 1)) (i : Fin n) (h : 0 < i) :

                                          Embedding a positive Fin n results in a positive Fin (n + 1)

                                          @[simp]
                                          theorem Fin.succAbove_castLT {n : } {x : Fin (n + 1)} {y : Fin (n + 1)} (h : x < y) (hx : optParam (x < n) (_ : x < n)) :
                                          @[simp]
                                          theorem Fin.succAbove_pred {n : } {x : Fin (n + 1)} {y : Fin (n + 1)} (h : x < y) (hy : optParam (y 0) (_ : y 0)) :
                                          theorem Fin.castLT_succAbove {n : } {x : Fin n} {y : Fin (n + 1)} (h : Fin.castSucc x < y) (h' : optParam (↑(Fin.succAbove y x) < n) (_ : ↑(Fin.succAbove y x) < n)) :
                                          theorem Fin.pred_succAbove {n : } {x : Fin n} {y : Fin (n + 1)} (h : y Fin.castSucc x) (h' : optParam (Fin.succAbove y x 0) (_ : Fin.succAbove y x 0)) :
                                          theorem Fin.exists_succAbove_eq {n : } {x : Fin (n + 1)} {y : Fin (n + 1)} (h : x y) :
                                          z, Fin.succAbove y z = x
                                          @[simp]
                                          theorem Fin.exists_succAbove_eq_iff {n : } {x : Fin (n + 1)} {y : Fin (n + 1)} :
                                          (z, Fin.succAbove x z = y) y x
                                          @[simp]
                                          theorem Fin.range_succAbove {n : } (p : Fin (n + 1)) :

                                          The range of p.succAbove is everything except p.

                                          @[simp]
                                          theorem Fin.range_succ (n : ) :
                                          Set.range Fin.succ = {0}
                                          @[simp]
                                          theorem Fin.exists_succ_eq_iff {n : } {x : Fin (n + 1)} :
                                          (y, Fin.succ y = x) x 0

                                          Given a fixed pivot x : Fin (n + 1), x.succAbove is injective

                                          theorem Fin.succAbove_right_inj {n : } {a : Fin n} {b : Fin n} {x : Fin (n + 1)} :

                                          Given a fixed pivot x : Fin (n + 1), x.succAbove is injective

                                          succAbove is injective at the pivot

                                          @[simp]
                                          theorem Fin.succAbove_left_inj {n : } {x : Fin (n + 1)} {y : Fin (n + 1)} :

                                          succAbove is injective at the pivot

                                          @[simp]
                                          theorem Fin.zero_succAbove {n : } (i : Fin n) :
                                          @[simp]
                                          theorem Fin.succ_succAbove_zero {n : } [NeZero n] (i : Fin n) :
                                          @[simp]
                                          theorem Fin.succ_succAbove_succ {n : } (i : Fin (n + 1)) (j : Fin n) :
                                          @[simp]
                                          theorem Fin.succ_succAbove_one {n : } [NeZero n] (i : Fin (n + 1)) :

                                          By moving succ to the outside of this expression, we create opportunities for further simplification using succAbove_zero or succ_succAbove_zero.

                                          @[simp]
                                          @[simp]
                                          def Fin.predAbove {n : } (p : Fin n) (i : Fin (n + 1)) :
                                          Fin n

                                          predAbove p i embeds i : Fin (n+1) into Fin n by subtracting one if p < i.

                                          Equations
                                          Instances For
                                            theorem Fin.predAbove_left_monotone {n : } (i : Fin (n + 1)) :
                                            Monotone fun p => Fin.predAbove p i
                                            def Fin.castPred {n : } (i : Fin (n + 2)) :
                                            Fin (n + 1)

                                            castPred embeds i : Fin (n + 2) into Fin (n + 1) by lowering just last (n + 1) to last n.

                                            Equations
                                            Instances For
                                              @[simp]
                                              theorem Fin.castPred_zero {n : } :
                                              @[simp]
                                              theorem Fin.castPred_one {n : } :
                                              @[simp]
                                              theorem Fin.predAbove_zero {n : } {i : Fin (n + 2)} (hi : i 0) :
                                              @[simp]
                                              theorem Fin.castPred_mk (n : ) (i : ) (h : i < n + 1) :
                                              Fin.castPred { val := i, isLt := (_ : i < Nat.succ (n + 1)) } = { val := i, isLt := h }
                                              @[simp]
                                              theorem Fin.castPred_mk' (n : ) (i : ) (h₁ : i < n + 2) (h₂ : i < n + 1) :
                                              Fin.castPred { val := i, isLt := h₁ } = { val := i, isLt := h₂ }
                                              theorem Fin.coe_castPred {n : } (a : Fin (n + 2)) (hx : a < Fin.last (n + 1)) :
                                              ↑(Fin.castPred a) = a
                                              theorem Fin.predAbove_below {n : } (p : Fin (n + 1)) (i : Fin (n + 2)) (h : i Fin.castSucc p) :
                                              @[simp]
                                              theorem Fin.predAbove_last {n : } :
                                              Fin.predAbove (Fin.last n) = Fin.castPred
                                              theorem Fin.predAbove_last_apply {n : } (i : Fin n) :
                                              Fin.predAbove (Fin.last n) i = Fin.castPred i
                                              theorem Fin.predAbove_above {n : } (p : Fin n) (i : Fin (n + 1)) (h : Fin.castSucc p < i) :
                                              Fin.predAbove p i = Fin.pred i (_ : i 0)
                                              theorem Fin.castPred_monotone {n : } :
                                              Monotone Fin.castPred
                                              @[simp]
                                              theorem Fin.succAbove_predAbove {n : } {p : Fin n} {i : Fin (n + 1)} (h : i Fin.castSucc p) :

                                              Sending Fin (n+1) to Fin n by subtracting one from anything above p then back to Fin (n+1) with a gap around p is the identity away from p.

                                              @[simp]
                                              theorem Fin.predAbove_succAbove {n : } (p : Fin n) (i : Fin n) :

                                              Sending Fin n into Fin (n + 1) with a gap at p then back to Fin n by subtracting one from anything above p is the identity.

                                              theorem Fin.castSucc_pred_eq_pred_castSucc {n : } {a : Fin (n + 1)} (ha : a 0) (ha' : optParam (Fin.castSucc a 0) (_ : Fin.castSucc a 0)) :
                                              theorem Fin.pred_succAbove_pred {n : } {a : Fin (n + 2)} {b : Fin (n + 1)} (ha : a 0) (hb : b 0) (hk : optParam (Fin.succAbove a b 0) (_ : Fin.succAbove a b 0)) :

                                              pred commutes with succAbove.

                                              @[simp]
                                              theorem Fin.succ_predAbove_succ {n : } (a : Fin n) (b : Fin (n + 1)) :

                                              succ commutes with predAbove.

                                              @[simp]
                                              theorem Fin.castPred_castSucc {n : } (i : Fin (n + 1)) :
                                              theorem Fin.castSucc_castPred {n : } {i : Fin (n + 2)} (h : i < Fin.last (n + 1)) :
                                              theorem Fin.coe_castPred_le_self {n : } (i : Fin (n + 2)) :
                                              ↑(Fin.castPred i) i
                                              theorem Fin.coe_castPred_lt_iff {n : } {i : Fin (n + 2)} :
                                              ↑(Fin.castPred i) < i i = Fin.last (n + 1)
                                              theorem Fin.lt_last_iff_coe_castPred {n : } {i : Fin (n + 2)} :
                                              i < Fin.last (n + 1) ↑(Fin.castPred i) = i
                                              @[simp]
                                              theorem Fin.coe_ofNat_eq_mod (m : ) (n : ) [NeZero m] :
                                              n = n % m

                                              mul #

                                              theorem Fin.mul_one' {n : } [NeZero n] (k : Fin n) :
                                              k * 1 = k
                                              theorem Fin.one_mul' {n : } [NeZero n] (k : Fin n) :
                                              1 * k = k
                                              theorem Fin.mul_zero' {n : } [NeZero n] (k : Fin n) :
                                              k * 0 = 0
                                              theorem Fin.zero_mul' {n : } [NeZero n] (k : Fin n) :
                                              0 * k = 0
                                              instance Fin.toExpr (n : ) :
                                              Equations
                                              • One or more equations did not get rendered due to their size.