General recursion in the SKI calculus #
In this file we implement general recursion functions (on the natural numbers), inspired by the
formalisation of Mathlib.Computability.Partrec. Since composition (B-combinator) and pairs
(MkPair, Fst, Snd) have been implemented in Cslib.Computability.CombinatoryLogic.Basic,
what remains are the following definitions and proofs of their correctness.
- Church numerals : a predicate
IsChurch n aexpressing that the termais βη-equivalent to the standard church numeraln— that is,a ⬝ f ⬝ x ↠ f ⬝ (f ⬝ ... ⬝ (f ⬝ x))). - SKI numerals :
ZeroandSucc, corresponding toPartrec.zeroandPartrec.succ, and correctness proofszero_correctandsucc_correct. - Predecessor : a term
Predso that (pred_correct)IsChurch n a → IsChurch n.pred (Pred ⬝ a). - Primitive recursion : a term
Recso that (rec_correct_succ)IsChurch (n+1) aimpliesRec ⬝ x ⬝ g ⬝ a ↠ g ⬝ a ⬝ (Rec ⬝ x ⬝ g ⬝ (Pred ⬝ a))and (rec_correct_zero)IsChurch 0 aimpliesRec ⬝ x ⬝ g ⬝ a ↠ x. - Unbounded root finding (μ-recursion) : given a term
frepresenting a functionfℕ: Nat → Nat, which takes on the value 0 a termRFindsuch that (rFind_correct)RFind ⬝ f ↠ asuch thatIsChurch n afornthe smallest root offℕ.
References #
- For church numerals and recursion via the fixed-point combinator, see sections 3.2 and 3.3 of Selinger's notes https://www.mscs.dal.ca/~selinger/papers/papers/lambdanotes.pdf
TODO #
- One could unify
is_bool,IsChurchandIsChurchPairinto a predicaterepresents : α → SKI → Prop, for any typeα"built from pieces that we understand" — something along the lines of "pure finite types" (see eg https://en.wikipedia.org/wiki/Primitive_recursive_functional). This would also clean up the statement ofrfind_correct. - The predicate
∃ n : Nat, IsChurch n : SKI → Propis semidecidable: by confluence, it suffices to normal-order reducea ⬝ f ⬝ xfor any "atomic" termsfandx. This could be implemented by defining reduction on polynomials. - With such a decision procedure, every SKI-term defines a partial function
Nat →. Nat, in the sense ofMathlib.Data.Part(as used inMathlib.Computability.Partrec). - The results of this file should define a surjection
SKI → Nat.Partrec.
Function form of the church numerals.
Equations
- Cslib.SKI.Church 0 f x = x
- Cslib.SKI.Church n_2.succ f x = f.app (Cslib.SKI.Church n_2 f x)
Instances For
The term a is βη-equivalent to a standard church numeral.
Equations
- Cslib.SKI.IsChurch n a = ∀ (f x : Cslib.SKI), Cslib.SKI.RedSKI.MRed ((a.app f).app x) (Cslib.SKI.Church n f x)
Instances For
Church numeral basics #
Church zero := λ f x. x
Equations
Instances For
Church one := λ f x. f x
Equations
Instances For
Church succ := λ a f x. f (a f x) λ a f. B f (a f) λ a. S B a ~ S B
Equations
Instances For
To define the predecessor, iterate the function PredAux ⟨i, j⟩ ↦ ⟨j, j+1⟩ on ⟨0,0⟩, then take
the first component.
Equations
- One or more equations did not get rendered due to their size.
Instances For
A term representing PredAux
Equations
Instances For
Useful auxiliary definition expressing that p represents ns ∈ Nat × Nat.
Equations
- Cslib.SKI.IsChurchPair ns x = (Cslib.SKI.IsChurch ns.fst (Cslib.SKI.Fst.app x) ∧ Cslib.SKI.IsChurch ns.snd (Cslib.SKI.Snd.app x))
Instances For
Predecessor := λ n. Fst ⬝ (n ⬝ PredAux ⬝ (MkPair ⬝ Zero ⬝ Zero))
Equations
- One or more equations did not get rendered due to their size.
Instances For
A term representing Pred
Equations
Instances For
Primitive recursion #
IsZero := λ n. n (K FF) TT
Equations
Instances For
A term representing IsZero
Equations
Instances For
To define Rec x g n := if n==0 then x else (Rec x g (Pred n)), we obtain a fixed point of
R ↦ λ x g n. Cond ⬝ (IsZero ⬝ n) ⬝ x ⬝ (g ⬝ a ⬝ (R ⬝ x ⬝ g ⬝ (Pred ⬝ n)))
Equations
- One or more equations did not get rendered due to their size.
Instances For
A term representing RecAux
Equations
Instances For
We define Rec so that
Rec ⬝ x ⬝ g ⬝ a ↠ SKI.Cond ⬝ x ⬝ (g ⬝ a ⬝ (Rec ⬝ x ⬝ g ⬝ (Pred ⬝ a))) ⬝ (IsZero ⬝ a)
Equations
Instances For
Root-finding (μ-recursion) #
First define an auxiliary function RFindAbove that looks for roots above a fixed number n, as a
fixed point of R ↦ λ n f. if f n = 0 then n else R f (n+1)
~ λ n f. Cond ⬝ n (R f (Succ n)) (IsZero (f n))
Equations
- One or more equations did not get rendered due to their size.
Instances For
A term representing RFindAboveAux
Instances For
Find the minimal root of fNat above a number n
Instances For
Ordinary root finding is root finding above zero
Instances For
Further numeric operations #
Addition: λ n m. n Succ m
Equations
Instances For
A term representing addition on church numerals
Equations
Instances For
Multiplication: λ n m. n (Add m) Zero
Equations
Instances For
A term representing multiplication on church numerals
Equations
Instances For
Subtraction: λ n m. n Pred m
Equations
Instances For
A term representing subtraction on church numerals
Equations
Instances For
Comparison: (. ≤ .) := λ n m. IsZero ⬝ (Sub ⬝ n ⬝ m)
Equations
Instances For
A term representing comparison on church numerals