Documentation

Mathlib.Order.SuccPred.LinearLocallyFinite

Linear locally finite orders #

We prove that a LinearOrder which is a LocallyFiniteOrder also verifies

Furthermore, we show that there is an OrderIso between such an order and a subset of .

Main definitions #

Main results #

Results about linear locally finite orders:

About toZ:

noncomputable def LinearLocallyFiniteOrder.succFn {ι : Type u_1} [LinearOrder ι] (i : ι) :
ι

Successor in a linear order. This defines a true successor only when i is isolated from above, i.e. when i is not the greatest lower bound of (i, ∞).

Equations
    Instances For
      theorem LinearLocallyFiniteOrder.isGLB_Ioc_of_isGLB_Ioi {ι : Type u_1} [LinearOrder ι] {i j k : ι} (hij_lt : i < j) (h : IsGLB (Set.Ioi i) k) :
      IsGLB (Set.Ioc i j) k
      theorem LinearLocallyFiniteOrder.succFn_le_of_lt {ι : Type u_1} [LinearOrder ι] (i j : ι) (hij : i < j) :
      theorem LinearLocallyFiniteOrder.le_of_lt_succFn {ι : Type u_1} [LinearOrder ι] (j i : ι) (hij : j < succFn i) :
      j i

      A locally finite order is a SuccOrder. This is not an instance, because its succ field conflicts with computable SuccOrder structures on and .

      Equations
        Instances For

          A locally finite order is a PredOrder. This is not an instance, because its succ field conflicts with computable PredOrder structures on and .

          Equations
            Instances For
              def toZ {ι : Type u_1} [LinearOrder ι] [SuccOrder ι] [IsSuccArchimedean ι] [PredOrder ι] (i0 i : ι) :

              toZ numbers elements of ι according to their order, starting from i0. We prove in orderIsoRangeToZOfLinearSuccPredArch that this defines an OrderIso between ι and the range of toZ.

              Equations
                Instances For
                  theorem toZ_of_ge {ι : Type u_1} [LinearOrder ι] [SuccOrder ι] [IsSuccArchimedean ι] [PredOrder ι] {i0 i : ι} (hi : i0 i) :
                  toZ i0 i = (Nat.find )
                  theorem toZ_of_lt {ι : Type u_1} [LinearOrder ι] [SuccOrder ι] [IsSuccArchimedean ι] [PredOrder ι] {i0 i : ι} (hi : i < i0) :
                  toZ i0 i = -(Nat.find )
                  @[simp]
                  theorem toZ_of_eq {ι : Type u_1} [LinearOrder ι] [SuccOrder ι] [IsSuccArchimedean ι] [PredOrder ι] {i0 : ι} :
                  toZ i0 i0 = 0
                  theorem iterate_succ_toZ {ι : Type u_1} [LinearOrder ι] [SuccOrder ι] [IsSuccArchimedean ι] [PredOrder ι] {i0 : ι} (i : ι) (hi : i0 i) :
                  theorem iterate_pred_toZ {ι : Type u_1} [LinearOrder ι] [SuccOrder ι] [IsSuccArchimedean ι] [PredOrder ι] {i0 : ι} (i : ι) (hi : i < i0) :
                  theorem toZ_nonneg {ι : Type u_1} [LinearOrder ι] [SuccOrder ι] [IsSuccArchimedean ι] [PredOrder ι] {i0 i : ι} (hi : i0 i) :
                  0 toZ i0 i
                  theorem toZ_neg {ι : Type u_1} [LinearOrder ι] [SuccOrder ι] [IsSuccArchimedean ι] [PredOrder ι] {i0 i : ι} (hi : i < i0) :
                  toZ i0 i < 0
                  theorem toZ_iterate_succ_le {ι : Type u_1} [LinearOrder ι] [SuccOrder ι] [IsSuccArchimedean ι] [PredOrder ι] {i0 : ι} (n : ) :
                  toZ i0 (Order.succ^[n] i0) n
                  theorem toZ_iterate_pred_ge {ι : Type u_1} [LinearOrder ι] [SuccOrder ι] [IsSuccArchimedean ι] [PredOrder ι] {i0 : ι} (n : ) :
                  -n toZ i0 (Order.pred^[n] i0)
                  theorem toZ_iterate_succ_of_not_isMax {ι : Type u_1} [LinearOrder ι] [SuccOrder ι] [IsSuccArchimedean ι] [PredOrder ι] {i0 : ι} (n : ) (hn : ¬IsMax (Order.succ^[n] i0)) :
                  toZ i0 (Order.succ^[n] i0) = n
                  theorem toZ_iterate_pred_of_not_isMin {ι : Type u_1} [LinearOrder ι] [SuccOrder ι] [IsSuccArchimedean ι] [PredOrder ι] {i0 : ι} (n : ) (hn : ¬IsMin (Order.pred^[n] i0)) :
                  toZ i0 (Order.pred^[n] i0) = -n
                  theorem le_of_toZ_le {ι : Type u_1} [LinearOrder ι] [SuccOrder ι] [IsSuccArchimedean ι] [PredOrder ι] {i0 i j : ι} (h_le : toZ i0 i toZ i0 j) :
                  i j
                  theorem toZ_mono {ι : Type u_1} [LinearOrder ι] [SuccOrder ι] [IsSuccArchimedean ι] [PredOrder ι] {i0 i j : ι} (h_le : i j) :
                  toZ i0 i toZ i0 j
                  theorem toZ_le_iff {ι : Type u_1} [LinearOrder ι] [SuccOrder ι] [IsSuccArchimedean ι] [PredOrder ι] {i0 : ι} (i j : ι) :
                  toZ i0 i toZ i0 j i j
                  theorem toZ_iterate_succ {ι : Type u_1} [LinearOrder ι] [SuccOrder ι] [IsSuccArchimedean ι] [PredOrder ι] {i0 : ι} [NoMaxOrder ι] (n : ) :
                  toZ i0 (Order.succ^[n] i0) = n
                  theorem toZ_iterate_pred {ι : Type u_1} [LinearOrder ι] [SuccOrder ι] [IsSuccArchimedean ι] [PredOrder ι] {i0 : ι} [NoMinOrder ι] (n : ) :
                  toZ i0 (Order.pred^[n] i0) = -n
                  theorem injective_toZ {ι : Type u_1} [LinearOrder ι] [SuccOrder ι] [IsSuccArchimedean ι] [PredOrder ι] {i0 : ι} :
                  noncomputable def orderIsoRangeToZOfLinearSuccPredArch {ι : Type u_1} [LinearOrder ι] [SuccOrder ι] [PredOrder ι] [IsSuccArchimedean ι] [ : Nonempty ι] :
                  ι ≃o (Set.range (toZ .some))

                  toZ defines an OrderIso between ι and its range.

                  Equations
                    Instances For
                      @[instance 100]
                      noncomputable def orderIsoIntOfLinearSuccPredArch {ι : Type u_1} [LinearOrder ι] [SuccOrder ι] [PredOrder ι] [IsSuccArchimedean ι] [NoMaxOrder ι] [NoMinOrder ι] [ : Nonempty ι] :

                      If the order has neither bot nor top, toZ defines an OrderIso between ι and .

                      Equations
                        Instances For

                          If the order has a bot but no top, toZ defines an OrderIso between ι and .

                          Equations
                            Instances For

                              If the order has both a bot and a top, toZ gives an OrderIso between ι and Finset.range n for some n.

                              Equations
                                Instances For