Documentation

Mathlib.Algebra.Order.AddGroupWithTop

Linearly ordered commutative additive groups and monoids with a top element adjoined #

This file sets up a special class of linearly ordered commutative additive monoids that show up as the target of so-called “valuations” in algebraic number theory.

Usually, in the informal literature, these objects are constructed by taking a linearly ordered commutative additive group Γ and formally adjoining a top element: Γ ∪ {⊤}.

The disadvantage is that a type such as ENNReal is not of that form, whereas it is a very common target for valuations. The solutions is to use a typeclass, and that is exactly what we do in this file.

A linearly ordered commutative monoid with an additively absorbing element. Instances should include number systems with an infinite element adjoined.

Instances

    A linearly ordered commutative group with an additively absorbing element. Instances should include number systems with an infinite element adjoined.

    Instances
      @[simp]
      theorem top_add {α : Type u_1} [LinearOrderedAddCommMonoidWithTop α] (a : α) :
      @[simp]
      theorem add_top {α : Type u_1} [LinearOrderedAddCommMonoidWithTop α] (a : α) :
      Equations
      • WithTop.LinearOrderedAddCommGroup.instNeg = { neg := Option.map fun (a : α) => -a }

      If α has subtraction, we can extend the subtraction to WithTop α, by setting x - ⊤ = ⊤ and ⊤ - x = ⊤. This definition is only registered as an instance on linearly ordered additive commutative groups, to avoid conflicting with the instance WithTop.instSub on types with a bottom element.

      Equations
      Instances For
        Equations
        • WithTop.LinearOrderedAddCommGroup.instSub = { sub := WithTop.LinearOrderedAddCommGroup.sub }
        @[simp]
        theorem WithTop.LinearOrderedAddCommGroup.coe_neg {α : Type u_1} [LinearOrderedAddCommGroup α] (a : α) :
        (-a) = -a
        @[simp]
        theorem WithTop.LinearOrderedAddCommGroup.coe_sub {α : Type u_1} [LinearOrderedAddCommGroup α] {a b : α} :
        (a - b) = a - b
        Equations
        @[instance 100]
        Equations