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 #
castAddMonoidHom
:cast
bundled as anAddMonoidHom
.castRingHom
:cast
bundled as aRingHom
.
Nat.cast : ℕ → α
as an AddMonoidHom
.
Equations
- Nat.castAddMonoidHom α = { toFun := Nat.cast, map_zero' := ⋯, map_add' := ⋯ }
Instances For
@[simp]
@[simp]
Equations
- Nat.castRingHom α = { toFun := Nat.cast, map_one' := ⋯, map_mul' := ⋯, map_zero' := ⋯, map_add' := ⋯ }
Instances For
@[simp]
@[simp]
Alias of Nat.cast_dvd_cast
.
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 : ℕ)
:
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 : ℕ)
:
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]
:
theorem
ext_nat''
{A : Type u_3}
{F : Type u_4}
[MulZeroOneClass A]
[FunLike F ℕ A]
[MonoidWithZeroHomClass F ℕ A]
(f g : F)
(h_pos : ∀ {n : ℕ}, 0 < n → f n = g n)
:
If two MonoidWithZeroHom
s agree on the positive naturals they are equal.
@[simp]
theorem
eq_natCast
{R : Type u_3}
{F : Type u_5}
[NonAssocSemiring R]
[FunLike F ℕ R]
[RingHomClass F ℕ R]
(f : F)
(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 : ℕ)
:
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]
:
This lemma is not marked @[simp]
lemma because its #discr_tree_key
(for the LHS) would just
be DFunLike.coe _ _
, due to the ofNat
that https://github.com/leanprover/lean4/issues/2867
forces us to include, and therefore it would negatively impact performance.
If that issue is resolved, this can be marked @[simp]
.
theorem
ext_nat
{R : Type u_3}
{F : Type u_5}
[NonAssocSemiring R]
[FunLike F ℕ R]
[RingHomClass F ℕ R]
(f g : F)
:
This is primed to match eq_intCast'
.
We don't use RingHomClass
here, since that might cause type-class slowdown for
Subsingleton
.
Equations
- Nat.uniqueRingHom = { default := Nat.castRingHom R, uniq := ⋯ }
instance
Pi.instNatCast
{α : Type u_1}
{π : α → Type u_3}
[(a : α) → NatCast (π a)]
:
NatCast ((a : α) → π a)
Equations
- Pi.instNatCast = { natCast := fun (n : ℕ) (x : α) => ↑n }
@[simp]
theorem
Pi.ofNat_apply
{α : Type u_1}
{π : α → Type u_3}
[(a : α) → NatCast (π a)]
(n : ℕ)
[n.AtLeastTwo]
(a : α)
:
theorem
Pi.ofNat_def
{α : Type u_1}
{π : α → Type u_3}
[(a : α) → NatCast (π a)]
(n : ℕ)
[n.AtLeastTwo]
: