Documentation

Mathlib.Tactic.NormNum.DivMod

norm_num extension for integer div/mod and divides #

This file adds support for the %, /, and (divisibility) operators on to the norm_num tactic.

theorem Mathlib.Meta.NormNum.isInt_ediv_zero {a b r : } :
IsInt a rIsNat b 0IsNat (a / b) 0
theorem Mathlib.Meta.NormNum.isInt_ediv {a b q m a' : } {b' r : } (ha : IsInt a a') (hb : IsNat b b') (hm : q * b' = m) (h : r + m = a') (h₂ : r.blt b' = true) :
IsInt (a / b) q
theorem Mathlib.Meta.NormNum.isInt_ediv_neg {a b q q' : } (h : IsInt (a / -b) q) (hq : -q = q') :
IsInt (a / b) q'

The norm_num extension which identifies expressions of the form Int.ediv a b, such that norm_num successfully recognises both a and b.

Equations
  • One or more equations did not get rendered due to their size.
Instances For
    def Mathlib.Meta.NormNum.evalIntDiv.core (a na : Q()) (za : ) (pa : Q(IsInt «$a» «$na»)) (b : Q()) (nb : Q()) (pb : Q(IsNat «$b» «$nb»)) :
    × (q : Q()) × Q(IsInt («$a» / «$b») «$q»)

    Given a result for evaluating a b in where b > 0, evaluate a / b.

    Instances For
      theorem Mathlib.Meta.NormNum.isInt_emod_zero {a b r : } :
      IsInt a rIsNat b 0IsInt (a % b) r
      theorem Mathlib.Meta.NormNum.isInt_emod {a b q m a' : } {b' r : } (ha : IsInt a a') (hb : IsNat b b') (hm : q * b' = m) (h : r + m = a') (h₂ : r.blt b' = true) :
      IsNat (a % b) r
      theorem Mathlib.Meta.NormNum.isInt_emod_neg {a b : } {r : } (h : IsNat (a % -b) r) :
      IsNat (a % b) r

      The norm_num extension which identifies expressions of the form Int.emod a b, such that norm_num successfully recognises both a and b.

      Equations
      • One or more equations did not get rendered due to their size.
      Instances For
        def Mathlib.Meta.NormNum.evalIntMod.go (a na : Q()) (za : ) (pa : Q(IsInt «$a» «$na»)) (b : Q()) :
        Result bOption (Result q(«$a» % «$b»))

        Given a result for evaluating a b in , evaluate a % b.

        Equations
        Instances For
          def Mathlib.Meta.NormNum.evalIntMod.core (a na : Q()) (za : ) (pa : Q(IsInt «$a» «$na»)) (b : Q()) (nb : Q()) (pb : Q(IsNat «$b» «$nb»)) :
          (r : Q()) × Q(IsNat («$a» % «$b») «$r»)

          Given a result for evaluating a b in where b > 0, evaluate a % b.

          Instances For
            theorem Mathlib.Meta.NormNum.isInt_dvd_true {a b a' b' c : } :
            IsInt a a'IsInt b b'a'.mul c = b'a b
            theorem Mathlib.Meta.NormNum.isInt_dvd_false {a b a' b' : } :
            IsInt a a'IsInt b b'(b'.emod a' != 0) = true¬a b

            The norm_num extension which identifies expressions of the form (a : ℤ) ∣ b, such that norm_num successfully recognises both a and b.

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