Encodings and Cardinality of First-Order Syntax #
Main Definitions #
- FirstOrder.Language.Term.encodingencodes terms as lists.
- FirstOrder.Language.BoundedFormula.encodingencodes bounded formulas as lists.
Main Results #
- FirstOrder.Language.Term.card_leshows that the number of terms in- L.Term αis at most- max ℵ₀ # (α ⊕ Σ i, L.Functions i).
- FirstOrder.Language.BoundedFormula.card_leshows that the number of bounded formulas in- Σ n, L.BoundedFormula α nis at most- max ℵ₀ (Cardinal.lift.{max u v} #α + Cardinal.lift.{u'} L.card).
TODO #
- Primcodableinstances for terms and formulas, based on the- encodings
- Computability facts about term and formula operations, to set up a computability approach to incompleteness
Encodes a term as a list of variables and function symbols.
Equations
- (FirstOrder.Language.var i).listEncode = [Sum.inl i]
- (FirstOrder.Language.func f ts).listEncode = Sum.inr ⟨l, f⟩ :: List.flatMap (fun (i : Fin l) => (ts i).listEncode) (List.finRange l)
Instances For
Decodes a list of variables and function symbols as a list of terms.
Equations
- One or more equations did not get rendered due to their size.
- FirstOrder.Language.Term.listDecode [] = []
- FirstOrder.Language.Term.listDecode (Sum.inl a :: l) = FirstOrder.Language.var a :: FirstOrder.Language.Term.listDecode l
Instances For
theorem
FirstOrder.Language.Term.listDecode_encode_list
{L : Language}
{α : Type u'}
(l : List (L.Term α))
 :
def
FirstOrder.Language.Term.encoding
{L : Language}
{α : Type u'}
 :
Computability.Encoding (L.Term α)
An encoding of terms as lists.
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[simp]
@[simp]
theorem
FirstOrder.Language.Term.card_sigma
{L : Language}
{α : Type u'}
 :
Cardinal.mk ((n : ℕ) × L.Term (α ⊕ Fin n)) = max Cardinal.aleph0 (Cardinal.mk (α ⊕ (i : ℕ) × L.Functions i))
instance
FirstOrder.Language.Term.small
{L : Language}
{α : Type u'}
[Small.{u, u'} α]
 :
Small.{u, max u u'} (L.Term α)
Encodes a bounded formula as a list of symbols.
Equations
- FirstOrder.Language.BoundedFormula.falsum.listEncode = [Sum.inr (Sum.inr (x✝ + 2))]
- (FirstOrder.Language.BoundedFormula.equal t₁ t₂).listEncode = [Sum.inl ⟨x✝, t₁⟩, Sum.inl ⟨x✝, t₂⟩]
- (FirstOrder.Language.BoundedFormula.rel R ts).listEncode = [Sum.inr (Sum.inl ⟨l, R⟩), Sum.inr (Sum.inr x✝)] ++ List.map (fun (i : Fin l) => Sum.inl ⟨x✝, ts i⟩) (List.finRange l)
- (φ₁.imp φ₂).listEncode = Sum.inr (Sum.inr 0) :: φ₁.listEncode ++ φ₂.listEncode
- φ.all.listEncode = Sum.inr (Sum.inr 1) :: φ.listEncode
Instances For
def
FirstOrder.Language.BoundedFormula.sigmaAll
{L : Language}
{α : Type u'}
 :
(n : ℕ) × L.BoundedFormula α n → (n : ℕ) × L.BoundedFormula α n
Applies the forall quantifier to an element of (Σ n, L.BoundedFormula α n),
or returns default if not possible.
Equations
Instances For
def
FirstOrder.Language.BoundedFormula.sigmaImp
{L : Language}
{α : Type u'}
 :
(n : ℕ) × L.BoundedFormula α n → (n : ℕ) × L.BoundedFormula α n → (n : ℕ) × L.BoundedFormula α n
Applies imp to two elements of (Σ n, L.BoundedFormula α n),
or returns default if not possible.
Equations
Instances For
@[irreducible]
Decodes a list of symbols as a list of formulas.
Equations
- One or more equations did not get rendered due to their size.
- FirstOrder.Language.BoundedFormula.listDecode (Sum.inr (Sum.inr n.succ.succ) :: l) = ⟨n, FirstOrder.Language.BoundedFormula.falsum⟩ :: FirstOrder.Language.BoundedFormula.listDecode l
- FirstOrder.Language.BoundedFormula.listDecode x✝ = []
Instances For
@[simp]
theorem
FirstOrder.Language.BoundedFormula.listDecode_encode_list
{L : Language}
{α : Type u'}
(l : List ((n : ℕ) × L.BoundedFormula α n))
 :
def
FirstOrder.Language.BoundedFormula.encoding
{L : Language}
{α : Type u'}
 :
Computability.Encoding ((n : ℕ) × L.BoundedFormula α n)
An encoding of bounded formulas as lists.
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[simp]
theorem
FirstOrder.Language.BoundedFormula.encoding_encode
{L : Language}
{α : Type u'}
(φ : (n : ℕ) × L.BoundedFormula α n)
 :
theorem
FirstOrder.Language.BoundedFormula.listEncode_sigma_injective
{L : Language}
{α : Type u'}
 :
Function.Injective fun (φ : (n : ℕ) × L.BoundedFormula α n) => φ.snd.listEncode
theorem
FirstOrder.Language.BoundedFormula.card_le
{L : Language}
{α : Type u'}
 :
Cardinal.mk ((n : ℕ) × L.BoundedFormula α n) ≤   max Cardinal.aleph0 (Cardinal.lift.{max u v, u'} (Cardinal.mk α) + Cardinal.lift.{u', max u v} L.card)