Documentation

Mathlib.Data.Nat.Cast.Basic

Cast of natural numbers (additional theorems) #

This file proves additional properties about the canonical homomorphism from the natural numbers into an additive monoid with a one (Nat.cast).

Main declarations #

Nat.cast : ℕ → α as an AddMonoidHom.

Equations
@[simp]
theorem Nat.coe_castAddMonoidHom {α : Type u_1} [AddMonoidWithOne α] :
(Nat.castAddMonoidHom α) = Nat.cast
theorem Even.natCast {α : Type u_1} [AddMonoidWithOne α] {n : } (hn : Even n) :
Even n
@[simp]
theorem Nat.cast_mul {α : Type u_1} [NonAssocSemiring α] (m : ) (n : ) :
(m * n) = m * n

Nat.cast : ℕ → α as a RingHom

Equations
  • Nat.castRingHom α = { toFun := Nat.cast, map_one' := , map_mul' := , map_zero' := , map_add' := }
@[simp]
theorem Nat.coe_castRingHom {α : Type u_1} [NonAssocSemiring α] :
(Nat.castRingHom α) = Nat.cast
theorem nsmul_eq_mul' {α : Type u_1} [NonAssocSemiring α] (a : α) (n : ) :
n a = a * n
@[simp]
theorem nsmul_eq_mul {α : Type u_1} [NonAssocSemiring α] (n : ) (a : α) :
n a = n * a
@[simp]
theorem Nat.cast_pow {α : Type u_1} [Semiring α] (m : ) (n : ) :
(m ^ n) = m ^ n
theorem Nat.cast_dvd_cast {α : Type u_1} [Semiring α] {m : } {n : } (h : m n) :
m n
theorem Dvd.dvd.natCast {α : Type u_1} [Semiring α] {m : } {n : } (h : m n) :
m n

Alias of Nat.cast_dvd_cast.

theorem ext_nat' {A : Type u_3} {F : Type u_5} [FunLike F A] [AddMonoid A] [AddMonoidHomClass F A] (f : F) (g : F) (h : f 1 = g 1) :
f = g
theorem AddMonoidHom.ext_nat_iff {A : Type u_3} [AddMonoid A] {f : →+ A} {g : →+ A} :
f = g f 1 = g 1
theorem AddMonoidHom.ext_nat {A : Type u_3} [AddMonoid A] {f : →+ A} {g : →+ A} :
f 1 = g 1f = g
theorem eq_natCast' {A : Type u_3} {F : Type u_5} [FunLike F A] [AddMonoidWithOne A] [AddMonoidHomClass F A] (f : F) (h1 : f 1 = 1) (n : ) :
f n = n
theorem map_natCast' {B : Type u_4} {F : Type u_5} [AddMonoidWithOne B] {A : Type u_6} [AddMonoidWithOne A] [FunLike F A B] [AddMonoidHomClass F A B] (f : F) (h : f 1 = 1) (n : ) :
f n = n
theorem map_ofNat' {B : Type u_4} {F : Type u_5} [AddMonoidWithOne B] {A : Type u_6} [AddMonoidWithOne A] [FunLike F A B] [AddMonoidHomClass F A B] (f : F) (h : f 1 = 1) (n : ) [n.AtLeastTwo] :
@[simp]
theorem nsmul_one {A : Type u_6} [AddMonoidWithOne A] (n : ) :
n 1 = n
theorem ext_nat'' {A : Type u_3} {F : Type u_4} [MulZeroOneClass A] [FunLike F A] [MonoidWithZeroHomClass F A] (f : F) (g : F) (h_pos : ∀ {n : }, 0 < nf n = g n) :
f = g

If two MonoidWithZeroHoms agree on the positive naturals they are equal.

theorem MonoidWithZeroHom.ext_nat_iff {A : Type u_3} [MulZeroOneClass A] {f : →*₀ A} {g : →*₀ A} :
f = g ∀ {n : }, 0 < nf n = g n
theorem MonoidWithZeroHom.ext_nat {A : Type u_3} [MulZeroOneClass A] {f : →*₀ A} {g : →*₀ A} :
(∀ {n : }, 0 < nf n = g n)f = g
@[simp]
theorem eq_natCast {R : Type u_3} {F : Type u_5} [NonAssocSemiring R] [FunLike F R] [RingHomClass F R] (f : F) (n : ) :
f n = n
@[simp]
theorem map_natCast {R : Type u_3} {S : Type u_4} {F : Type u_5} [NonAssocSemiring R] [NonAssocSemiring S] [FunLike F R S] [RingHomClass F R S] (f : F) (n : ) :
f n = n
theorem map_ofNat {R : Type u_3} {S : Type u_4} {F : Type u_5} [NonAssocSemiring R] [NonAssocSemiring S] [FunLike F R S] [RingHomClass F R S] (f : F) (n : ) [n.AtLeastTwo] :
theorem ext_nat {R : Type u_3} {F : Type u_5} [NonAssocSemiring R] [FunLike F R] [RingHomClass F R] (f : F) (g : F) :
f = g
theorem NeZero.nat_of_neZero {R : Type u_6} {S : Type u_7} [Semiring R] [Semiring S] {F : Type u_8} [FunLike F R S] [RingHomClass F R S] (f : F) {n : } [hn : NeZero n] :
NeZero n

This is primed to match eq_intCast'.

@[simp]
theorem Nat.cast_id (n : ) :
n = n

We don't use RingHomClass here, since that might cause type-class slowdown for Subsingleton

Equations
def multiplesHom (β : Type u_2) [AddMonoid β] :
β ( →+ β)

Additive homomorphisms from are defined by the image of 1.

Equations
  • multiplesHom β = { toFun := fun (x : β) => { toFun := fun (n : ) => n x, map_zero' := , map_add' := }, invFun := fun (f : →+ β) => f 1, left_inv := , right_inv := }
def powersHom (α : Type u_1) [Monoid α] :

Monoid homomorphisms from Multiplicative are defined by the image of Multiplicative.ofAdd 1.

Equations
@[simp]
theorem multiplesHom_apply (β : Type u_2) [AddMonoid β] (x : β) (n : ) :
((multiplesHom β) x) n = n x
@[simp]
theorem powersHom_apply {α : Type u_1} [Monoid α] (x : α) (n : Multiplicative ) :
((powersHom α) x) n = x ^ Multiplicative.toAdd n
@[simp]
theorem multiplesHom_symm_apply (β : Type u_2) [AddMonoid β] (f : →+ β) :
(multiplesHom β).symm f = f 1
@[simp]
theorem powersHom_symm_apply {α : Type u_1} [Monoid α] (f : Multiplicative →* α) :
(powersHom α).symm f = f (Multiplicative.ofAdd 1)
theorem MonoidHom.apply_mnat {α : Type u_1} [Monoid α] (f : Multiplicative →* α) (n : Multiplicative ) :
f n = f (Multiplicative.ofAdd 1) ^ Multiplicative.toAdd n
theorem MonoidHom.ext_mnat_iff {α : Type u_1} [Monoid α] {f : Multiplicative →* α} {g : Multiplicative →* α} :
f = g f (Multiplicative.ofAdd 1) = g (Multiplicative.ofAdd 1)
theorem MonoidHom.ext_mnat {α : Type u_1} [Monoid α] ⦃f : Multiplicative →* α ⦃g : Multiplicative →* α (h : f (Multiplicative.ofAdd 1) = g (Multiplicative.ofAdd 1)) :
f = g
theorem AddMonoidHom.apply_nat (β : Type u_2) [AddMonoid β] (f : →+ β) (n : ) :
f n = n f 1
def multiplesAddHom (β : Type u_2) [AddCommMonoid β] :
β ≃+ ( →+ β)

If α is commutative, multiplesHom is an additive equivalence.

Equations
def powersMulHom (α : Type u_1) [CommMonoid α] :

If α is commutative, powersHom is a multiplicative equivalence.

Equations
@[simp]
theorem multiplesAddHom_apply (β : Type u_2) [AddCommMonoid β] (x : β) (n : ) :
((multiplesAddHom β) x) n = n x
@[simp]
theorem powersMulHom_apply (α : Type u_1) [CommMonoid α] (x : α) (n : Multiplicative ) :
((powersMulHom α) x) n = x ^ Multiplicative.toAdd n
@[simp]
theorem multiplesAddHom_symm_apply (β : Type u_2) [AddCommMonoid β] (f : →+ β) :
(multiplesAddHom β).symm f = f 1
@[simp]
theorem powersMulHom_symm_apply (α : Type u_1) [CommMonoid α] (f : Multiplicative →* α) :
(powersMulHom α).symm f = f (Multiplicative.ofAdd 1)
instance Pi.instNatCast {α : Type u_1} {π : αType u_3} [(a : α) → NatCast (π a)] :
NatCast ((a : α) → π a)
Equations
  • Pi.instNatCast = { natCast := fun (n : ) (x : α) => n }
theorem Pi.natCast_apply {α : Type u_1} {π : αType u_3} [(a : α) → NatCast (π a)] (n : ) (a : α) :
n a = n
@[simp]
theorem Pi.natCast_def {α : Type u_1} {π : αType u_3} [(a : α) → NatCast (π a)] (n : ) :
n = fun (x : α) => n
@[deprecated Pi.natCast_apply]
theorem Pi.nat_apply {α : Type u_1} {π : αType u_3} [(a : α) → NatCast (π a)] (n : ) (a : α) :
n a = n

Alias of Pi.natCast_apply.

@[deprecated Pi.natCast_def]
theorem Pi.coe_nat {α : Type u_1} {π : αType u_3} [(a : α) → NatCast (π a)] (n : ) :
n = fun (x : α) => n

Alias of Pi.natCast_def.

@[simp]
theorem Pi.ofNat_apply {α : Type u_1} {π : αType u_3} [(a : α) → NatCast (π a)] (n : ) [n.AtLeastTwo] (a : α) :
OfNat.ofNat n a = n
theorem Sum.elim_natCast_natCast {α : Type u_3} {β : Type u_4} {γ : Type u_5} [NatCast γ] (n : ) :
Sum.elim n n = n