Theory Goodstein_Lambda
section ‹Specification›
theory Goodstein_Lambda
imports Main
begin
subsection ‹Hereditary base representation›
text ‹We define a data type of trees and an evaluation function that sums siblings and
exponentiates with respect to the given base on nesting.›
datatype C = C (unC: "C list")
fun evalC where
"evalC b (C []) = 0"
| "evalC b (C (x # xs)) = b^evalC b x + evalC b (C xs)"
value "evalC 2 (C [])"
value "evalC 2 (C [C []])"
value "evalC 2 (C [C [C []]])"
value "evalC 2 (C [C [], C []])"
text ‹The hereditary base representation is characterized as trees (i.e., nested lists) whose
lists have monotonically increasing evaluations, with fewer than @{term "b"} repetitions for
each value. We will show later that this representation is unique.›
inductive_set hbase for b where
"C [] ∈ hbase b"
| "i ≠ 0 ⟹ i < b ⟹ n ∈ hbase b ⟹
C ms ∈ hbase b ⟹ (⋀m'. m' ∈ set ms ⟹ evalC b n < evalC b m') ⟹
C (replicate i n @ ms) ∈ hbase b"
text ‹We can convert to and from natural numbers as follows.›
definition H2N where
"H2N b n = evalC b n"
text ‹As we will show later, @{term "H2N b"} restricted to @{term "hbase n"} is bijective
if @{prop "b ≥ (2 :: nat)"}, so we can convert from natural numbers by taking the inverse.›
definition N2H where
"N2H b n = inv_into (hbase b) (H2N b) n"
subsection ‹The Goodstein function›
text ‹We define a function that computes the length of the Goodstein sequence whose $c$-th element
is $g_c = n$. Termination will be shown later, thereby establishing Goodstein's theorem.›
function (sequential) goodstein :: "nat ⇒ nat ⇒ nat" where
"goodstein 0 n = 0"
| "goodstein c 0 = c"
| "goodstein c n = goodstein (c+1) (H2N (c+2) (N2H (c+1) n) - 1)"
by pat_completeness auto
abbreviation 𝒢 where
"𝒢 n ≡ goodstein (Suc 0) n"
section ‹Ordinals›
text ‹The following type contains countable ordinals, by the usual case distinction into 0,
successor ordinal, or limit ordinal; limit ordinals are given by their fundamental sequence.
Hereditary base @{term "b"} representations carry over to such ordinals by replacing each
occurrence of the base by @{term "ω"}.›