Definitions of basic functions #
theorem
Int.subNatNat_of_sub_eq_succ
{m n k : Nat}
(h : n - m = k.succ)
:
Int.subNatNat m n = Int.negSucc k
These are only for internal use #
theorem
Int.negSucc_add_negSucc
(m n : Nat)
:
Int.negSucc m + Int.negSucc n = Int.negSucc (m + n).succ
theorem
Int.negSucc_mul_negSucc'
(m n : Nat)
:
Int.negSucc m * Int.negSucc n = Int.ofNat (m.succ * n.succ)
theorem
Int.subNatNat_elim
(m n : Nat)
(motive : Nat → Nat → Int → Prop)
(hp : ∀ (i n : Nat), motive (n + i) n ↑i)
(hn : ∀ (i m : Nat), motive m (m + i + 1) (Int.negSucc i))
:
motive m n (Int.subNatNat m n)
Equations
Equations
theorem
Int.ofNat_add_negSucc_of_lt
{m n : Nat}
(h : m < n.succ)
:
Int.ofNat m + Int.negSucc n = Int.negSucc (n - m)
theorem
Int.subNatNat_sub
{n m : Nat}
(h : n ≤ m)
(k : Nat)
:
Int.subNatNat (m - n) k = Int.subNatNat m (k + n)
theorem
Int.subNatNat_add_negSucc
(m n k : Nat)
:
Int.subNatNat m n + Int.negSucc k = Int.subNatNat m (n + k.succ)
theorem
Int.add_assoc.aux2
(m n k : Nat)
:
Int.negSucc m + Int.negSucc n + ↑k = Int.negSucc m + (Int.negSucc n + ↑k)
Equations
@[simp]
Equations
theorem
Int.negSucc_mul_negOfNat
(m n : Nat)
:
Int.negSucc m * Int.negOfNat n = Int.ofNat (m.succ * n)
theorem
Int.negOfNat_mul_negSucc
(m n : Nat)
:
Int.negOfNat n * Int.negSucc m = Int.ofNat (n * m.succ)
Equations
theorem
Int.ofNat_mul_subNatNat
(m n k : Nat)
:
↑m * Int.subNatNat n k = Int.subNatNat (m * n) (m * k)
theorem
Int.negSucc_mul_subNatNat
(m n k : Nat)
:
Int.negSucc m * Int.subNatNat n k = Int.subNatNat (m.succ * k) (m.succ * n)
Equations
NatCast lemmas
The following lemmas are later subsumed by e.g. Nat.cast_add
and Nat.cast_mul
in Mathlib
but it is convenient to have these earlier, for users who only need Nat
and Int
.