Documentation

Mathlib.ModelTheory.LanguageMap

Language Maps #

Maps between first-order languages in the style of the Flypitch project, as well as several important maps between structures.

Main Definitions #

References #

For the Flypitch project:

structure FirstOrder.Language.LHom (L : Language) (L' : Language) :
Type (max (max (max u u') v) v')

A language homomorphism maps the symbols of one language to symbols of another.

Instances For

    A language homomorphism maps the symbols of one language to symbols of another.

    Equations
    • One or more equations did not get rendered due to their size.
    Instances For
      def FirstOrder.Language.LHom.reduct {L : Language} {L' : Language} (ϕ : L →ᴸ L') (M : Type u_1) [L'.Structure M] :

      Pulls a structure back along a language map.

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

        The identity language homomorphism.

        Equations
        Instances For
          @[simp]
          @[simp]

          The inclusion of the left factor into the sum of two languages.

          Equations
          Instances For

            The inclusion of the right factor into the sum of two languages.

            Equations
            Instances For

              The inclusion of an empty language into any other language.

              Equations
              Instances For
                theorem FirstOrder.Language.LHom.funext {L : Language} {L' : Language} {F G : L →ᴸ L'} (h_fun : F.onFunction = G.onFunction) (h_rel : F.onRelation = G.onRelation) :
                F = G
                def FirstOrder.Language.LHom.comp {L : Language} {L' : Language} {L'' : Language} (g : L' →ᴸ L'') (f : L →ᴸ L') :
                L →ᴸ L''

                The composition of two language homomorphisms.

                Equations
                Instances For
                  @[simp]
                  theorem FirstOrder.Language.LHom.comp_onFunction {L : Language} {L' : Language} {L'' : Language} (g : L' →ᴸ L'') (f : L →ᴸ L') (_n : ) (F : L.Functions _n) :
                  @[simp]
                  theorem FirstOrder.Language.LHom.comp_onRelation {L : Language} {L' : Language} {L'' : Language} (g : L' →ᴸ L'') (f : L →ᴸ L') (x✝ : ) (R : L.Relations x✝) :
                  @[simp]
                  theorem FirstOrder.Language.LHom.id_comp {L : Language} {L' : Language} (F : L →ᴸ L') :
                  (LHom.id L').comp F = F
                  @[simp]
                  theorem FirstOrder.Language.LHom.comp_id {L : Language} {L' : Language} (F : L →ᴸ L') :
                  F.comp (LHom.id L) = F
                  theorem FirstOrder.Language.LHom.comp_assoc {L : Language} {L' : Language} {L'' : Language} {L3 : Language} (F : L'' →ᴸ L3) (G : L' →ᴸ L'') (H : L →ᴸ L') :
                  (F.comp G).comp H = F.comp (G.comp H)
                  def FirstOrder.Language.LHom.sumElim {L : Language} {L' : Language} (ϕ : L →ᴸ L') {L'' : Language} (ψ : L'' →ᴸ L') :
                  L.sum L'' →ᴸ L'

                  A language map defined on two factors of a sum.

                  Equations
                  • One or more equations did not get rendered due to their size.
                  Instances For
                    @[simp]
                    theorem FirstOrder.Language.LHom.sumElim_onFunction {L : Language} {L' : Language} (ϕ : L →ᴸ L') {L'' : Language} (ψ : L'' →ᴸ L') (_n : ) (a✝ : L.Functions _n L''.Functions _n) :
                    (ϕ.sumElim ψ).onFunction a✝ = Sum.elim (fun (f : L.Functions _n) => ϕ.onFunction f) (fun (f : L''.Functions _n) => ψ.onFunction f) a✝
                    @[simp]
                    theorem FirstOrder.Language.LHom.sumElim_onRelation {L : Language} {L' : Language} (ϕ : L →ᴸ L') {L'' : Language} (ψ : L'' →ᴸ L') (_n : ) (a✝ : L.Relations _n L''.Relations _n) :
                    (ϕ.sumElim ψ).onRelation a✝ = Sum.elim (fun (f : L.Relations _n) => ϕ.onRelation f) (fun (f : L''.Relations _n) => ψ.onRelation f) a✝
                    theorem FirstOrder.Language.LHom.sumElim_comp_inl {L : Language} {L' : Language} (ϕ : L →ᴸ L') {L'' : Language} (ψ : L'' →ᴸ L') :
                    theorem FirstOrder.Language.LHom.sumElim_comp_inr {L : Language} {L' : Language} (ϕ : L →ᴸ L') {L'' : Language} (ψ : L'' →ᴸ L') :
                    theorem FirstOrder.Language.LHom.comp_sumElim {L : Language} {L' : Language} (ϕ : L →ᴸ L') {L'' : Language} (ψ : L'' →ᴸ L') {L3 : Language} (θ : L' →ᴸ L3) :
                    θ.comp (ϕ.sumElim ψ) = (θ.comp ϕ).sumElim (θ.comp ψ)
                    def FirstOrder.Language.LHom.sumMap {L : Language} {L' : Language} (ϕ : L →ᴸ L') {L₁ : Language} {L₂ : Language} (ψ : L₁ →ᴸ L₂) :
                    L.sum L₁ →ᴸ L'.sum L₂

                    The map between two sum-languages induced by maps on the two factors.

                    Equations
                    • One or more equations did not get rendered due to their size.
                    Instances For
                      @[simp]
                      theorem FirstOrder.Language.LHom.sumMap_onRelation {L : Language} {L' : Language} (ϕ : L →ᴸ L') {L₁ : Language} {L₂ : Language} (ψ : L₁ →ᴸ L₂) (_n : ) (a✝ : L.Relations _n L₁.Relations _n) :
                      (ϕ.sumMap ψ).onRelation a✝ = Sum.map (fun (f : L.Relations _n) => ϕ.onRelation f) (fun (f : L₁.Relations _n) => ψ.onRelation f) a✝
                      @[simp]
                      theorem FirstOrder.Language.LHom.sumMap_onFunction {L : Language} {L' : Language} (ϕ : L →ᴸ L') {L₁ : Language} {L₂ : Language} (ψ : L₁ →ᴸ L₂) (_n : ) (a✝ : L.Functions _n L₁.Functions _n) :
                      (ϕ.sumMap ψ).onFunction a✝ = Sum.map (fun (f : L.Functions _n) => ϕ.onFunction f) (fun (f : L₁.Functions _n) => ψ.onFunction f) a✝
                      @[simp]
                      theorem FirstOrder.Language.LHom.sumMap_comp_inl {L : Language} {L' : Language} (ϕ : L →ᴸ L') {L₁ : Language} {L₂ : Language} (ψ : L₁ →ᴸ L₂) :
                      @[simp]
                      theorem FirstOrder.Language.LHom.sumMap_comp_inr {L : Language} {L' : Language} (ϕ : L →ᴸ L') {L₁ : Language} {L₂ : Language} (ψ : L₁ →ᴸ L₂) :
                      structure FirstOrder.Language.LHom.Injective {L : Language} {L' : Language} (ϕ : L →ᴸ L') :

                      A language homomorphism is injective when all the maps between symbol types are.

                      Instances For
                        noncomputable def FirstOrder.Language.LHom.defaultExpansion {L : Language} {L' : Language} (ϕ : L →ᴸ L') [(n : ) → (f : L'.Functions n) → Decidable (f Set.range fun (f : L.Functions n) => ϕ.onFunction f)] [(n : ) → (r : L'.Relations n) → Decidable (r Set.range fun (r : L.Relations n) => ϕ.onRelation r)] (M : Type u_1) [Inhabited M] [L.Structure M] :

                        Pulls an L-structure along a language map ϕ : L →ᴸ L', and then expands it to an L'-structure arbitrarily.

                        Equations
                        • One or more equations did not get rendered due to their size.
                        Instances For
                          class FirstOrder.Language.LHom.IsExpansionOn {L : Language} {L' : Language} (ϕ : L →ᴸ L') (M : Type u_1) [L.Structure M] [L'.Structure M] :

                          A language homomorphism is an expansion on a structure if it commutes with the interpretation of all symbols on that structure.

                          Instances
                            @[simp]
                            theorem FirstOrder.Language.LHom.map_onFunction {L : Language} {L' : Language} (ϕ : L →ᴸ L') {M : Type u_1} [L.Structure M] [L'.Structure M] [ϕ.IsExpansionOn M] {n : } (f : L.Functions n) (x : Fin nM) :
                            @[simp]
                            theorem FirstOrder.Language.LHom.map_onRelation {L : Language} {L' : Language} (ϕ : L →ᴸ L') {M : Type u_1} [L.Structure M] [L'.Structure M] [ϕ.IsExpansionOn M] {n : } (R : L.Relations n) (x : Fin nM) :
                            instance FirstOrder.Language.LHom.sumElim_isExpansionOn {L : Language} {L' : Language} (ϕ : L →ᴸ L') {L'' : Language} (ψ : L'' →ᴸ L') (M : Type u_1) [L.Structure M] [L'.Structure M] [L''.Structure M] [ϕ.IsExpansionOn M] [ψ.IsExpansionOn M] :
                            instance FirstOrder.Language.LHom.sumMap_isExpansionOn {L : Language} {L' : Language} (ϕ : L →ᴸ L') {L₁ : Language} {L₂ : Language} (ψ : L₁ →ᴸ L₂) (M : Type u_1) [L.Structure M] [L'.Structure M] [L₁.Structure M] [L₂.Structure M] [ϕ.IsExpansionOn M] [ψ.IsExpansionOn M] :
                            @[simp]
                            theorem FirstOrder.Language.LHom.funMap_sumInl {L : Language} {L' : Language} {M : Type w} [L.Structure M] [(L.sum L').Structure M] [LHom.sumInl.IsExpansionOn M] {n : } {f : L.Functions n} {x : Fin nM} :
                            @[simp]
                            theorem FirstOrder.Language.LHom.funMap_sumInr {L : Language} {L' : Language} {M : Type w} [L.Structure M] [(L'.sum L).Structure M] [LHom.sumInr.IsExpansionOn M] {n : } {f : L.Functions n} {x : Fin nM} :
                            @[instance 100]
                            instance FirstOrder.Language.LHom.isExpansionOn_reduct {L : Language} {L' : Language} (ϕ : L →ᴸ L') (M : Type u_1) [L'.Structure M] :
                            theorem FirstOrder.Language.LHom.Injective.isExpansionOn_default {L : Language} {L' : Language} {ϕ : L →ᴸ L'} [(n : ) → (f : L'.Functions n) → Decidable (f Set.range fun (f : L.Functions n) => ϕ.onFunction f)] [(n : ) → (r : L'.Relations n) → Decidable (r Set.range fun (r : L.Relations n) => ϕ.onRelation r)] (h : ϕ.Injective) (M : Type u_1) [Inhabited M] [L.Structure M] :
                            structure FirstOrder.Language.LEquiv (L : Language) (L' : Language) :
                            Type (max (max (max u_1 u_2) u_3) u_4)

                            A language equivalence maps the symbols of one language to symbols of another bijectively.

                            Instances For

                              A language equivalence maps the symbols of one language to symbols of another bijectively.

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

                                The identity equivalence from a first-order language to itself.

                                Equations
                                Instances For

                                  The inverse of an equivalence of first-order languages.

                                  Equations
                                  Instances For
                                    def FirstOrder.Language.LEquiv.trans {L : Language} {L' : Language} {L'' : Language} (e : L ≃ᴸ L') (e' : L' ≃ᴸ L'') :
                                    L ≃ᴸ L''

                                    The composition of equivalences of first-order languages.

                                    Equations
                                    Instances For
                                      @[simp]
                                      theorem FirstOrder.Language.LEquiv.trans_invLHom {L : Language} {L' : Language} {L'' : Language} (e : L ≃ᴸ L') (e' : L' ≃ᴸ L'') :
                                      @[simp]
                                      theorem FirstOrder.Language.LEquiv.trans_toLHom {L : Language} {L' : Language} {L'' : Language} (e : L ≃ᴸ L') (e' : L' ≃ᴸ L'') :

                                      The type of functions for a language consisting only of constant symbols.

                                      Equations
                                      Instances For

                                        A language with constants indexed by a type.

                                        Equations
                                        Instances For
                                          @[simp]
                                          def FirstOrder.Language.constantsOn.structure {M : Type w} {α : Type u'} (f : αM) :

                                          Gives a constantsOn α structure to a type by assigning each constant a value.

                                          Equations
                                          • One or more equations did not get rendered due to their size.
                                          Instances For
                                            def FirstOrder.Language.LHom.constantsOnMap {α : Type u'} {β : Type v'} (f : αβ) :

                                            A map between index types induces a map between constant languages.

                                            Equations
                                            • One or more equations did not get rendered due to their size.
                                            Instances For
                                              theorem FirstOrder.Language.constantsOnMap_isExpansionOn {M : Type w} {α : Type u'} {β : Type v'} {f : αβ} {fα : αM} {fβ : βM} (h : f = ) :

                                              Extends a language with a constant for each element of a parameter set in M.

                                              Equations
                                              Instances For

                                                Extends a language with a constant for each element of a parameter set in M.

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

                                                  The language map adding constants.

                                                  Equations
                                                  Instances For
                                                    @[simp]
                                                    @[simp]
                                                    def FirstOrder.Language.con (L : Language) {α : Type w'} (a : α) :

                                                    The constant symbol indexed by a particular element.

                                                    Equations
                                                    Instances For

                                                      The language map removing an empty constant set.

                                                      Equations
                                                      • One or more equations did not get rendered due to their size.
                                                      Instances For
                                                        @[simp]
                                                        @[simp]
                                                        @[deprecated FirstOrder.Language.withConstants_funMap_sumInl (since := "2025-02-21")]

                                                        Alias of FirstOrder.Language.withConstants_funMap_sumInl.

                                                        @[deprecated FirstOrder.Language.withConstants_relMap_sumInl (since := "2025-02-21")]

                                                        Alias of FirstOrder.Language.withConstants_relMap_sumInl.

                                                        def FirstOrder.Language.lhomWithConstantsMap (L : Language) {α : Type w'} {β : Type u_1} (f : αβ) :

                                                        The language map extending the constant set.

                                                        Equations
                                                        Instances For
                                                          @[simp]
                                                          theorem FirstOrder.Language.withConstants_funMap_sumInr (L : Language) {M : Type w} [L.Structure M] (α : Type u_1) [(constantsOn α).Structure M] {a : α} {x : Fin 0M} :
                                                          Structure.funMap (Sum.inr a) x = (L.con a)
                                                          @[deprecated FirstOrder.Language.withConstants_funMap_sumInr (since := "2025-02-21")]
                                                          theorem FirstOrder.Language.withConstants_funMap_sum_inr (L : Language) {M : Type w} [L.Structure M] (α : Type u_1) [(constantsOn α).Structure M] {a : α} {x : Fin 0M} :
                                                          Structure.funMap (Sum.inr a) x = (L.con a)

                                                          Alias of FirstOrder.Language.withConstants_funMap_sumInr.

                                                          @[simp]
                                                          theorem FirstOrder.Language.coe_con (L : Language) {M : Type w} [L.Structure M] (A : Set M) {a : A} :
                                                          (L.con a) = a