Documentation

Mathlib.Algebra.Order.Ring.StandardPart

Standard part function #

Given a finite element in a non-archimedean field, the standard part function rounds it to the unique closest real number. That is, it chops off any infinitesimals.

Let K be a linearly ordered field. The subset of finite elements (i.e. those bounded by a natural number) is a ValuationSubring, which means we can construct its residue field FiniteResidueField, roughly corresponding to the finite elements quotiented by infinitesimals. This field inherits a LinearOrder instance, which makes it into an Archimedean linearly ordered field, meaning we can uniquely embed it in the reals.

Given a finite element of the field, the ArchimedeanClass.stdPart function returns the real number corresponding to this unique embedding. This function generalizes, among other things, the standard part function on Hyperreal.

TODO #

Redefine Hyperreal.st in terms of ArchimedeanClass.stdPart.

References #

Finite residue field #

noncomputable def ArchimedeanClass.FiniteElement (K : Type u_1) [LinearOrder K] [Field K] [IsOrderedRing K] :
Type u_1

The valuation subring of elements in non-negative Archimedean classes, i.e. elements bounded by some natural number.

Equations
Instances For
    @[simp]
    theorem ArchimedeanClass.FiniteElement.val_add {K : Type u_1} [LinearOrder K] [Field K] [IsOrderedRing K] (x y : FiniteElement K) :
    ↑(x + y) = x + y
    @[simp]
    theorem ArchimedeanClass.FiniteElement.val_sub {K : Type u_1} [LinearOrder K] [Field K] [IsOrderedRing K] (x y : FiniteElement K) :
    ↑(x - y) = x - y
    @[simp]
    theorem ArchimedeanClass.FiniteElement.val_mul {K : Type u_1} [LinearOrder K] [Field K] [IsOrderedRing K] (x y : FiniteElement K) :
    ↑(x * y) = x * y
    theorem ArchimedeanClass.FiniteElement.ext {K : Type u_1} [LinearOrder K] [Field K] [IsOrderedRing K] {x y : FiniteElement K} (h : x = y) :
    x = y

    The constructor for FiniteElement.

    Equations
    Instances For
      @[simp]
      theorem ArchimedeanClass.FiniteElement.mk_natCast {K : Type u_1} [LinearOrder K] [Field K] [IsOrderedRing K] {n : } (h : 0 mk n) :
      FiniteElement.mk (↑n) h = n
      @[simp]
      theorem ArchimedeanClass.FiniteElement.mk_intCast {K : Type u_1} [LinearOrder K] [Field K] [IsOrderedRing K] {n : } (h : 0 mk n) :
      FiniteElement.mk (↑n) h = n
      @[simp]

      The residue field of FiniteElement. This quotient inherits an order from K, which makes it into a linearly ordered Archimedean field.

      Equations
      Instances For

        The quotient map from finite elements on the field to the associated residue field.

        Equations
        Instances For
          theorem ArchimedeanClass.FiniteResidueField.ind {K : Type u_1} [LinearOrder K] [Field K] [IsOrderedRing K] {motive : FiniteResidueField KProp} (mk : ∀ (x : FiniteElement K), motive (mk x)) (x : FiniteResidueField K) :
          motive x

          An embedding from an Archimedean field into K induces an embedding into FiniteResidueField K.

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

            Standard part #

            noncomputable def ArchimedeanClass.stdPart {K : Type u_1} [LinearOrder K] [Field K] [IsOrderedRing K] (x : K) :

            The standard part of a FiniteElement is the unique real number with an infinitesimal difference.

            For any infinite inputs, this function outputs a junk value of 0.

            Equations
            Instances For
              theorem ArchimedeanClass.stdPart_of_mk_ne_zero {K : Type u_1} [LinearOrder K] [Field K] [IsOrderedRing K] {x : K} (h : mk x 0) :
              @[simp]
              theorem ArchimedeanClass.stdPart_neg {K : Type u_1} [LinearOrder K] [Field K] [IsOrderedRing K] (x : K) :
              theorem ArchimedeanClass.stdPart_add {K : Type u_1} [LinearOrder K] [Field K] [IsOrderedRing K] {x y : K} (hx : 0 mk x) (hy : 0 mk y) :
              theorem ArchimedeanClass.stdPart_sub {K : Type u_1} [LinearOrder K] [Field K] [IsOrderedRing K] {x y : K} (hx : 0 mk x) (hy : 0 mk y) :
              theorem ArchimedeanClass.stdPart_mul {K : Type u_1} [LinearOrder K] [Field K] [IsOrderedRing K] {x y : K} (hx : 0 mk x) (hy : 0 mk y) :
              theorem ArchimedeanClass.stdPart_div {K : Type u_1} [LinearOrder K] [Field K] [IsOrderedRing K] {x y : K} (hx : 0 mk x) (hy : 0 -mk y) :
              @[simp]
              theorem ArchimedeanClass.stdPart_intCast {K : Type u_1} [LinearOrder K] [Field K] [IsOrderedRing K] (n : ) :
              stdPart n = n
              @[simp]
              theorem ArchimedeanClass.stdPart_natCast {K : Type u_1} [LinearOrder K] [Field K] [IsOrderedRing K] (n : ) :
              stdPart n = n
              @[simp]
              theorem ArchimedeanClass.stdPart_ratCast {K : Type u_1} [LinearOrder K] [Field K] [IsOrderedRing K] (q : ) :
              stdPart q = q
              @[simp]
              theorem ArchimedeanClass.stdPart_real {K : Type u_1} [LinearOrder K] [Field K] [IsOrderedRing K] (f : →+*o K) (r : ) :
              stdPart (f r) = r
              theorem ArchimedeanClass.mk_sub_pos_iff {K : Type u_1} [LinearOrder K] [Field K] [IsOrderedRing K] {x : K} (f : →+*o K) {r : } (hx : 0 mk x) :
              0 < mk (x - f r) stdPart x = r

              The standard part of x is the unique real r such that x - r is infinitesimal.

              theorem ArchimedeanClass.mk_sub_stdPart_pos {K : Type u_1} [LinearOrder K] [Field K] [IsOrderedRing K] {x : K} (f : →+*o K) (hx : 0 mk x) :
              0 < mk (x - f (stdPart x))