Theory

Exponential

It is proved that the graph of exponential is definable by -formula, and their inductive properties are provable in . In , we can further prove their entireness.

Theory of Hereditary Finite Sets in

Weak theory of sets in (Hereditary Finite Sets) can be developed inside using Ackermann coding and bit predicate. Hereafter, we will use the notation in the sense of bit predicate:

lemma mem_iff_bit [M ⊧ₘ* 𝐈𝚺₁] {i a : M} : i ∈ a ↔ Bit i a

The following comprehension holds.

theorem finset_comprehension₁ [M ⊧ₘ* 𝐈𝚺₁] {P : M → Prop} (hP : (Γ, 1)-Predicate P) (a : M) : ∃ s < exp a, ∀ i < a, i ∈ s ↔ P i

The basic concepts of set theory, such as union, inter, cartesian product, and mapping, etc. are defined.

Seq

iff is a mapping and its domain is for some .

def Seq [M ⊧ₘ* 𝐈𝚺₁] (s : M) : Prop := IsMapping s ∧ ∃ l, domain s = under l

Primitive Recursion

Let , be a -function. There is a -function such that:

structure Formulae (k : ℕ) where zero : 𝚺₁-Semisentence (k + 1) succ : 𝚺₁-Semisentence (k + 3) structure Construction {k : ℕ} (p : Formulae k) where zero : (Fin k → M) → M succ : (Fin k → M) → M → M → M zero_defined : DefinedFunction zero p.zero succ_defined : DefinedFunction (fun v ↦ succ (v ·.succ.succ) (v 1) (v 0)) p.succ variable {k : ℕ} {p : Formulae k} (c : Construction M p) (v : Fin k → M) def Construction.result (u : M) : M theorem Construction.result_zero : c.result v 0 = c.zero v theorem Construction.result_succ (u : M) : c.result v (u + 1) = c.succ v u (c.result v u)

Fixpoint

Let be a predicate, which takes a class as a parameter. Then there is a -predicate such that

if satisfies following conditions:

  1. is -definable if is a set. i.e., a predicate is -definable.
  2. Monotone: and implies .
  3. Finite: implies the existence of a s.t. .
structure Blueprint (k : ℕ) where core : 𝚫₁-Semisentence (k + 2) structure Construction (φ : Blueprint k) where Φ : (Fin k → M) → Set M → M → Prop defined : Arith.Defined (fun v ↦ Φ (v ·.succ.succ) {x | x ∈ v 1} (v 0)) φ.core monotone {C C' : Set M} (h : C ⊆ C') {v x} : Φ v C x → Φ v C' x class Construction.Finite (c : Construction M φ) where finite {C : Set M} {v x} : c.Φ v C x → ∃ m, c.Φ v {y ∈ C | y < m} x variable {k : ℕ} {φ : Blueprint k} (c : Construction M φ) [Finite c] (v : Fin k → M) def Construction.Fixpoint (x : M) : Prop theorem Construction.case : c.Fixpoint v x ↔ c.Φ v {z | c.Fixpoint v z} x

is if satisfies strong finiteness:

class Construction.StrongFinite (c : Construction M φ) where strong_finite {C : Set M} {v x} : c.Φ v C x → c.Φ v {y ∈ C | y < x} x

Also structural induction holds.

theorem Construction.induction [c.StrongFinite] {P : M → Prop} (hP : (Γ, 1)-Predicate P) (H : ∀ C : Set M, (∀ x ∈ C, c.Fixpoint v x ∧ P x) → ∀ x, c.Φ v C x → P x) : ∀ x, c.Fixpoint v x → P x