Documentation

Mathlib.ModelTheory.Order

Ordered First-Ordered Structures #

This file defines ordered first-order languages and structures, as well as their theories.

Main Definitions #

Main Results #

The type of relations for the language of orders, consisting of a single binary relation le.

Instances For

    The relational language consisting of a single relation representing .

    Equations
    Instances For
      @[simp]

      A language is ordered if it has a symbol representing .

      • leSymb : L.Relations 2

        The relation symbol representing .

      Instances
        def FirstOrder.Language.Term.le {L : Language} {α : Type w} {n : } [L.IsOrdered] (t₁ t₂ : L.Term (α Fin n)) :

        Joins two terms t₁, t₂ in a formula representing t₁ ≤ t₂.

        Equations
        Instances For
          def FirstOrder.Language.Term.lt {L : Language} {α : Type w} {n : } [L.IsOrdered] (t₁ t₂ : L.Term (α Fin n)) :

          Joins two terms t₁, t₂ in a formula representing t₁ < t₂.

          Equations
          Instances For

            The language homomorphism sending the unique symbol of Language.order to in an ordered language.

            Equations
            • One or more equations did not get rendered due to their size.
            Instances For
              @[simp]
              theorem FirstOrder.Language.orderLHom_onRelation (L : Language) [L.IsOrdered] (x✝ : ) (x✝¹ : Language.order.Relations x✝) :
              L.orderLHom.onRelation x✝¹ = match x✝, x✝¹ with | .(2), orderRel.le => leSymb

              A sentence indicating that an order has no top element: $\forall x, \exists y, \neg y \le x$.

              Equations
              Instances For

                A sentence indicating that an order has no bottom element: $\forall x, \exists y, \neg x \le y$.

                Equations
                Instances For

                  A sentence indicating that an order is dense: $\forall x, \forall y, x < y \to \exists z, x < z \wedge z < y$.

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

                    The theory of dense linear orders without endpoints.

                    Equations
                    Instances For

                      Any linearly-ordered type is naturally a structure in the language Language.order. This is not an instance, because sometimes the Language.order.Structure is defined first.

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

                        A structure is ordered if its language has a symbol whose interpretation is .

                        Instances
                          @[simp]
                          theorem FirstOrder.Language.Term.realize_le {L : Language} {α : Type w} {M : Type w'} {n : } [L.IsOrdered] [L.Structure M] [LE M] [L.OrderedStructure M] {t₁ t₂ : L.Term (α Fin n)} {v : αM} {xs : Fin nM} :
                          (t₁.le t₂).Realize v xs realize (Sum.elim v xs) t₁ realize (Sum.elim v xs) t₂
                          @[simp]
                          theorem FirstOrder.Language.Term.realize_lt {L : Language} {α : Type w} {M : Type w'} {n : } [L.IsOrdered] [L.Structure M] [Preorder M] [L.OrderedStructure M] {t₁ t₂ : L.Term (α Fin n)} {v : αM} {xs : Fin nM} :
                          (t₁.lt t₂).Realize v xs realize (Sum.elim v xs) t₁ < realize (Sum.elim v xs) t₂

                          Any structure in an ordered language can be ordered correspondingly.

                          Equations
                          Instances For
                            def FirstOrder.Language.decidableLEOfStructure (L : Language) (M : Type w') [L.IsOrdered] [L.Structure M] [h : DecidableRel fun (a b : M) => Structure.RelMap leSymb ![a, b]] :
                            DecidableRel fun (x1 x2 : M) => x1 x2

                            The order structure on an ordered language is decidable.

                            Equations
                            Instances For

                              Any model of a theory of preorders is a preorder.

                              Equations
                              Instances For

                                Any model of a theory of partial orders is a partial order.

                                Equations
                                Instances For

                                  Any model of a theory of linear orders is a linear order.

                                  Equations
                                  Instances For
                                    theorem FirstOrder.Language.HomClass.monotone {L : Language} {M : Type w'} [L.IsOrdered] [L.Structure M] {N : Type u_1} [L.Structure N] {F : Type u_2} [FunLike F M N] [L.HomClass F M N] [Preorder M] [L.OrderedStructure M] [Preorder N] [L.OrderedStructure N] (f : F) :
                                    theorem FirstOrder.Language.HomClass.strictMono {L : Language} {M : Type w'} [L.IsOrdered] [L.Structure M] {N : Type u_1} [L.Structure N] {F : Type u_2} [FunLike F M N] [L.HomClass F M N] [EmbeddingLike F M N] [PartialOrder M] [L.OrderedStructure M] [PartialOrder N] [L.OrderedStructure N] (f : F) :
                                    theorem FirstOrder.Language.StrongHomClass.toOrderIsoClass (L : Language) [L.IsOrdered] (M : Type u_1) [L.Structure M] [LE M] [L.OrderedStructure M] (N : Type u_2) [L.Structure N] [LE N] [L.OrderedStructure N] (F : Type u_3) [EquivLike F M N] [L.StrongHomClass F M N] :

                                    This is not an instance because it would form a loop with FirstOrder.Language.order.instStrongHomClassOfOrderIsoClass. As both types are Props, it would only cause a slowdown.

                                    Any countable nonempty model of the theory of dense linear orders is a Fraïssé limit of the class of finite models of the theory of linear orders.

                                    The theory of dense linear orders is ℵ₀-categorical.

                                    The theory of dense linear orders is ℵ₀-complete.