Finitely Generated First-Order Structures #
This file defines what it means for a first-order (sub)structure to be finitely or countably generated, similarly to other finitely-generated objects in the algebra library.
Main Definitions #
FirstOrder.Language.Substructure.FG
indicates that a substructure is finitely generated.FirstOrder.Language.Structure.FG
indicates that a structure is finitely generated.FirstOrder.Language.Substructure.CG
indicates that a substructure is countably generated.FirstOrder.Language.Structure.CG
indicates that a structure is countably generated.
TODO #
Develop a more unified definition of finite generation using the theory of closure operators, or use this definition of finite generation to define the others.
def
FirstOrder.Language.Substructure.FG
{L : Language}
{M : Type u_1}
[L.Structure M]
(N : L.Substructure M)
:
A substructure of M
is finitely generated if it is the closure of a finite subset of M
.
Instances For
instance
FirstOrder.Language.Substructure.instInhabited_fg
{L : Language}
{M : Type u_1}
[L.Structure M]
:
Equations
- FirstOrder.Language.Substructure.instInhabited_fg = { default := ⟨⊥, ⋯⟩ }
theorem
FirstOrder.Language.Substructure.FG.sup
{L : Language}
{M : Type u_1}
[L.Structure M]
{N₁ N₂ : L.Substructure M}
(hN₁ : N₁.FG)
(hN₂ : N₂.FG)
:
theorem
FirstOrder.Language.Substructure.FG.map
{L : Language}
{M : Type u_1}
[L.Structure M]
{N : Type u_2}
[L.Structure N]
(f : L.Hom M N)
{s : L.Substructure M}
(hs : s.FG)
:
(Substructure.map f s).FG
theorem
FirstOrder.Language.Substructure.FG.of_map_embedding
{L : Language}
{M : Type u_1}
[L.Structure M]
{N : Type u_2}
[L.Structure N]
(f : L.Embedding M N)
{s : L.Substructure M}
(hs : (Substructure.map f.toHom s).FG)
:
s.FG
theorem
FirstOrder.Language.Substructure.FG.of_finite
{L : Language}
{M : Type u_1}
[L.Structure M]
{s : L.Substructure M}
[h : Finite ↥s]
:
s.FG
theorem
FirstOrder.Language.Substructure.FG.finite
{L : Language}
{M : Type u_1}
[L.Structure M]
[L.IsRelational]
{S : L.Substructure M}
(h : S.FG)
:
Finite ↥S
theorem
FirstOrder.Language.Substructure.fg_iff_finite
{L : Language}
{M : Type u_1}
[L.Structure M]
[L.IsRelational]
{S : L.Substructure M}
:
def
FirstOrder.Language.Substructure.CG
{L : Language}
{M : Type u_1}
[L.Structure M]
(N : L.Substructure M)
:
A substructure of M
is countably generated if it is the closure of a countable subset of M
.
Equations
Instances For
theorem
FirstOrder.Language.Substructure.FG.cg
{L : Language}
{M : Type u_1}
[L.Structure M]
{N : L.Substructure M}
(h : N.FG)
:
N.CG
theorem
FirstOrder.Language.Substructure.CG.sup
{L : Language}
{M : Type u_1}
[L.Structure M]
{N₁ N₂ : L.Substructure M}
(hN₁ : N₁.CG)
(hN₂ : N₂.CG)
:
theorem
FirstOrder.Language.Substructure.CG.map
{L : Language}
{M : Type u_1}
[L.Structure M]
{N : Type u_2}
[L.Structure N]
(f : L.Hom M N)
{s : L.Substructure M}
(hs : s.CG)
:
(Substructure.map f s).CG
theorem
FirstOrder.Language.Substructure.CG.of_map_embedding
{L : Language}
{M : Type u_1}
[L.Structure M]
{N : Type u_2}
[L.Structure N]
(f : L.Embedding M N)
{s : L.Substructure M}
(hs : (Substructure.map f.toHom s).CG)
:
s.CG
theorem
FirstOrder.Language.Substructure.cg_of_countable
{L : Language}
{M : Type u_1}
[L.Structure M]
{s : L.Substructure M}
[h : Countable ↥s]
:
s.CG
An equivalent expression of Structure.FG
in terms of Set.Finite
instead of Finset
.
theorem
FirstOrder.Language.Structure.FG.finite
{L : Language}
{M : Type u_1}
[L.Structure M]
[L.IsRelational]
(h : FG L M)
:
Finite M
theorem
FirstOrder.Language.Structure.fg_iff_finite
{L : Language}
{M : Type u_1}
[L.Structure M]
[L.IsRelational]
:
theorem
FirstOrder.Language.Substructure.fg_iff_structure_fg
{L : Language}
{M : Type u_1}
[L.Structure M]
(S : L.Substructure M)
:
theorem
FirstOrder.Language.Substructure.cg_iff_structure_cg
{L : Language}
{M : Type u_1}
[L.Structure M]
(S : L.Substructure M)
: