Documentation

Mathlib.Tactic.NormNum.Result

The Result type for norm_num #

We set up predicates IsNat, IsInt, and IsRat, stating that an element of a ring is equal to the "normal form" of a natural number, integer, or rational number coerced into that ring.

We then define Result e, which contains a proof that a typed expression e : Q($α) is equal to the coercion of an explicit natural number, integer, or rational number, or is either true or false.

A shortcut (non)instance for AddMonoidWithOne to shrink generated proofs.

Equations
Instances For

    A shortcut (non)instance for AddMonoidWithOne α from Ring α to shrink generated proofs.

    Equations
    • Mathlib.Meta.NormNum.instAddMonoidWithOne = inferInstance
    Instances For

      Helper function to synthesize a typed AddMonoidWithOne α expression.

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

        Helper function to synthesize a typed Semiring α expression.

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

          Helper function to synthesize a typed Ring α expression.

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

            Represent an integer as a "raw" typed expression.

            This uses .lit (.natVal n) internally to represent a natural number, rather than the preferred OfNat.ofNat form. We use this internally to avoid unnecessary typeclass searches.

            This function is the inverse of Expr.intLit!.

            Instances For

              Represent an integer as a "raw" typed expression.

              This .lit (.natVal n) internally to represent a natural number, rather than the preferred OfNat.ofNat form. We use this internally to avoid unnecessary typeclass searches.

              Instances For
                def Mathlib.Meta.NormNum.rawIntLitNatAbs (n : Q()) :
                (m : Q()) × Q(Int.natAbs «$n» = «$m»)

                Extract the raw natlit representing the absolute value of a raw integer literal (of the type produced by Mathlib.Meta.NormNum.mkRawIntLit) along with an equality proof.

                Equations
                • One or more equations did not get rendered due to their size.
                Instances For
                  def Mathlib.Meta.NormNum.mkOfNat {u : Lean.Level} (α : Q(Type u)) (_sα : Q(AddMonoidWithOne «$α»)) (lit : Q()) :
                  Lean.MetaM ((a' : Q(«$α»)) × Q(«$lit» = «$a'»))

                  Constructs an ofNat application a' with the canonical instance, together with a proof that the instance is equal to the result of Nat.cast on the given AddMonoidWithOne instance.

                  This function is performance-critical, as many higher level tactics have to construct numerals. So rather than using typeclass search we hardcode the (relatively small) set of solutions to the typeclass problem.

                  Instances For
                    structure Mathlib.Meta.NormNum.IsNat {α : Type u_1} [AddMonoidWithOne α] (a : α) (n : ) :

                    Assert that an element of a semiring is equal to the coercion of some natural number.

                    • out : a = n

                      The element is equal to the coercion of the natural number.

                    Instances For
                      def Nat.rawCast {α : Type u_1} [AddMonoidWithOne α] (n : ) :
                      α

                      A "raw nat cast" is an expression of the form (Nat.rawCast lit : α) where lit is a raw natural number literal. These expressions are used by tactics like ring to decrease the number of typeclass arguments required in each use of a number literal at type α.

                      Equations
                      Instances For
                        theorem Mathlib.Meta.NormNum.IsNat.to_eq {α : Type u_1} [AddMonoidWithOne α] {n : } {a : α} {a' : α} :
                        Mathlib.Meta.NormNum.IsNat a nn = a'a = a'
                        theorem Mathlib.Meta.NormNum.isNat.natElim {p : Prop} {n : } {n' : } :
                        Mathlib.Meta.NormNum.IsNat n n'p n'p n
                        structure Mathlib.Meta.NormNum.IsInt {α : Type u_1} [Ring α] (a : α) (n : ) :

                        Assert that an element of a ring is equal to the coercion of some integer.

                        • out : a = n

                          The element is equal to the coercion of the integer.

                        Instances For
                          def Int.rawCast {α : Type u_1} [Ring α] (n : ) :
                          α

                          A "raw int cast" is an expression of the form:

                          • (Nat.rawCast lit : α) where lit is a raw natural number literal
                          • (Int.rawCast (Int.negOfNat lit) : α) where lit is a nonzero raw natural number literal

                          (That is, we only actually use this function for negative integers.) This representation is used by tactics like ring to decrease the number of typeclass arguments required in each use of a number literal at type α.

                          Equations
                          Instances For
                            theorem Mathlib.Meta.NormNum.IsInt.neg_to_eq {α : Type u_1} [Ring α] {n : } {a : α} {a' : α} :
                            Mathlib.Meta.NormNum.IsInt a (Int.negOfNat n)n = a'a = -a'
                            theorem Mathlib.Meta.NormNum.IsInt.nonneg_to_eq {α : Type u_1} [Ring α] {n : } {a : α} {a' : α} (h : Mathlib.Meta.NormNum.IsInt a (Int.ofNat n)) (e : n = a') :
                            a = a'
                            inductive Mathlib.Meta.NormNum.IsRat {α : Type u_1} [Ring α] (a : α) (num : ) (denom : ) :

                            Assert that an element of a ring is equal to num / denom (and denom is invertible so that this makes sense). We will usually also have num and denom coprime, although this is not part of the definition.

                            Instances For
                              def Rat.rawCast {α : Type u_1} [DivisionRing α] (n : ) (d : ) :
                              α

                              A "raw rat cast" is an expression of the form:

                              • (Nat.rawCast lit : α) where lit is a raw natural number literal
                              • (Int.rawCast (Int.negOfNat lit) : α) where lit is a nonzero raw natural number literal
                              • (Rat.rawCast n d : α) where n is a raw int literal, d is a raw nat literal, and d is not 1 or 0.

                              (where a raw int literal is of the form Int.ofNat lit or Int.negOfNat nzlit where lit is a raw nat literal)

                              This representation is used by tactics like ring to decrease the number of typeclass arguments required in each use of a number literal at type α.

                              Equations
                              Instances For
                                theorem Mathlib.Meta.NormNum.IsRat.to_raw_eq {α : Type u_1} {n : } {d : } [DivisionRing α] {a : α} :
                                theorem Mathlib.Meta.NormNum.IsRat.neg_to_eq {α : Type u_1} [DivisionRing α] {n : } {d : } {a : α} {n' : α} {d' : α} :
                                Mathlib.Meta.NormNum.IsRat a (Int.negOfNat n) dn = n'd = d'a = -(n' / d')
                                theorem Mathlib.Meta.NormNum.IsRat.nonneg_to_eq {α : Type u_1} [DivisionRing α] {n : } {d : } {a : α} {n' : α} {d' : α} :
                                Mathlib.Meta.NormNum.IsRat a (Int.ofNat n) dn = n'd = d'a = n' / d'
                                theorem Mathlib.Meta.NormNum.IsRat.den_nz {α : Type u_1} [DivisionRing α] {a : α} {n : } {d : } :

                                The result of norm_num running on an expression x of type α. Untyped version of Result.

                                Instances For
                                  def Mathlib.Meta.NormNum.Result {u : Lean.Level} {α : Q(Type u)} (x : Q(«$α»)) :

                                  The result of norm_num running on an expression x of type α.

                                  Equations
                                  Instances For
                                    instance Mathlib.Meta.NormNum.instInhabitedResult :
                                    {a : Lean.Level} → {α : Q(Type a)} → {x : Q(«$α»)} → Inhabited (Mathlib.Meta.NormNum.Result x)
                                    Equations
                                    @[match_pattern, inline]

                                    The result is proof : x, where x is a (true) proposition.

                                    Equations
                                    Instances For
                                      @[match_pattern, inline]

                                      The result is proof : ¬x, where x is a (false) proposition.

                                      Equations
                                      Instances For
                                        @[match_pattern, inline]
                                        def Mathlib.Meta.NormNum.Result.isNat {u : Lean.Level} {α : Q(Type u)} {x : Q(«$α»)} (inst : autoParam Q(AddMonoidWithOne «$α») _auto✝) (lit : Q()) (proof : Q(Mathlib.Meta.NormNum.IsNat «$x» «$lit»)) :

                                        The result is lit : ℕ (a raw nat literal) and proof : isNat x lit.

                                        Equations
                                        Instances For
                                          @[match_pattern, inline]
                                          def Mathlib.Meta.NormNum.Result.isNegNat {u : Lean.Level} {α : Q(Type u)} {x : Q(«$α»)} (inst : autoParam Q(Ring «$α») _auto✝) (lit : Q()) (proof : Q(Mathlib.Meta.NormNum.IsInt «$x» (Int.negOfNat «$lit»))) :

                                          The result is -lit where lit is a raw nat literal and proof : isInt x (.negOfNat lit).

                                          Equations
                                          Instances For
                                            @[match_pattern, inline]
                                            def Mathlib.Meta.NormNum.Result.isRat {u : Lean.Level} {α : Q(Type u)} {x : Q(«$α»)} (inst : autoParam Q(DivisionRing «$α») _auto✝) (q : ) (n : Q()) (d : Q()) (proof : Q(Mathlib.Meta.NormNum.IsRat «$x» «$n» «$d»)) :

                                            The result is proof : isRat x n d, where n is either .ofNat lit or .negOfNat lit with lit a raw nat literal and d is a raw nat literal (not 0 or 1), and q is the value of n / d.

                                            Equations
                                            Instances For
                                              def Mathlib.Meta.NormNum.Result.isInt {u : Lean.Level} {α : Q(Type u)} {x : Q(«$α»)} (inst : autoParam Q(Ring «$α») _auto✝) (z : Q()) (n : ) (proof : Q(Mathlib.Meta.NormNum.IsInt «$x» «$z»)) :

                                              The result is z : ℤ and proof : isNat x z.

                                              Equations
                                              • One or more equations did not get rendered due to their size.
                                              Instances For
                                                def Mathlib.Meta.NormNum.Result.isRat' {u : Lean.Level} {α : Q(Type u)} {x : Q(«$α»)} (inst : autoParam Q(DivisionRing «$α») _auto✝) (q : ) (n : Q()) (d : Q()) (proof : Q(Mathlib.Meta.NormNum.IsRat «$x» «$n» «$d»)) :

                                                The result depends on whether q : ℚ happens to be an integer, in which case the result is .isInt .. whereas otherwise it's .isRat ...

                                                Equations
                                                • One or more equations did not get rendered due to their size.
                                                Instances For
                                                  Equations
                                                  • One or more equations did not get rendered due to their size.
                                                  def Mathlib.Meta.NormNum.Result.toRat :
                                                  {a : Lean.Level} → {α : Q(Type a)} → {e : Q(«$α»)} → Mathlib.Meta.NormNum.Result eOption

                                                  Returns the rational number that is the result of norm_num evaluation.

                                                  Equations
                                                  • One or more equations did not get rendered due to their size.
                                                  Instances For
                                                    def Mathlib.Meta.NormNum.Result.toRatNZ :
                                                    {a : Lean.Level} → {α : Q(Type a)} → {e : Q(«$α»)} → Mathlib.Meta.NormNum.Result eOption ( × Option Lean.Expr)

                                                    Returns the rational number that is the result of norm_num evaluation, along with a proof that the denominator is nonzero in the isRat case.

                                                    Equations
                                                    • One or more equations did not get rendered due to their size.
                                                    Instances For
                                                      def Mathlib.Meta.NormNum.Result.toInt {u : Lean.Level} {α : Q(Type u)} {e : Q(«$α»)} (_i : autoParam Q(Ring «$α») _auto✝) :
                                                      Mathlib.Meta.NormNum.Result eOption ( × (lit : Q()) × Q(Mathlib.Meta.NormNum.IsInt «$e» «$lit»))

                                                      Extract from a Result the integer value (as both a term and an expression), and the proof that the original expression is equal to this integer.

                                                      Equations
                                                      • One or more equations did not get rendered due to their size.
                                                      Instances For
                                                        def Mathlib.Meta.NormNum.Result.toRat' {u : Lean.Level} {α : Q(Type u)} {e : Q(«$α»)} (_i : autoParam Q(DivisionRing «$α») _auto✝) :
                                                        Mathlib.Meta.NormNum.Result eOption ( × (n : Q()) × (d : Q()) × Q(Mathlib.Meta.NormNum.IsRat «$e» «$n» «$d»))

                                                        Extract from a Result the rational value (as both a term and an expression), and the proof that the original expression is equal to this rational number.

                                                        Equations
                                                        • One or more equations did not get rendered due to their size.
                                                        Instances For
                                                          def Mathlib.Meta.NormNum.Result.toRawEq {u : Lean.Level} {α : Q(Type u)} {e : Q(«$α»)} :
                                                          Mathlib.Meta.NormNum.Result e(e' : Q(«$α»)) × Q(«$e» = «$e'»)

                                                          Given a NormNum.Result e (which uses IsNat, IsInt, IsRat to express equality to a rational numeral), converts it to an equality e = Nat.rawCast n, e = Int.rawCast n, or e = Rat.rawCast n d to a raw cast expression, so it can be used for rewriting.

                                                          Equations
                                                          • One or more equations did not get rendered due to their size.
                                                          Instances For
                                                            def Mathlib.Meta.NormNum.Result.toRawIntEq {u : Lean.Level} {α : Q(Type u)} {e : Q(«$α»)} :
                                                            Mathlib.Meta.NormNum.Result eOption ( × (e' : Q(«$α»)) × Q(«$e» = «$e'»))

                                                            Result.toRawEq but providing an integer. Given a NormNum.Result e for something known to be an integer (which uses IsNat or IsInt to express equality to an integer numeral), converts it to an equality e = Nat.rawCast n or e = Int.rawCast n to a raw cast expression, so it can be used for rewriting. Gives none if not an integer.

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

                                                              Constructs a Result out of a raw nat cast. Assumes e is a raw nat cast expression.

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

                                                                Constructs a Result out of a raw int cast. Assumes e is a raw int cast expression denoting n.

                                                                Equations
                                                                • One or more equations did not get rendered due to their size.
                                                                Instances For
                                                                  def Mathlib.Meta.NormNum.Result.ofRawRat {u : Lean.Level} {α : Q(Type u)} (q : ) (e : Q(«$α»)) (hyp : optParam (Option Lean.Expr) none) :

                                                                  Constructs a Result out of a raw rat cast. Assumes e is a raw rat cast expression denoting n.

                                                                  Instances For

                                                                    Convert a Result to a Simp.Result.

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

                                                                      Given Mathlib.Meta.NormNum.Result.isBool p b, this is the type of p. Note that BoolResult p b is definitionally equal to Expr, and if you write match b with ..., then in the true branch BoolResult p true is reducibly equal to Q($p) and in the false branch it is reducibly equal to Q(¬ $p).

                                                                      Equations
                                                                      Instances For
                                                                        def Mathlib.Meta.NormNum.Result.eqTrans {u : Lean.Level} {α : Q(Type u)} {a : Q(«$α»)} {b : Q(«$α»)} (eq : Q(«$a» = «$b»)) :

                                                                        If a = b and we can evaluate b, then we can evaluate a.

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