The partial recursive functions #
The partial recursive functions are defined similarly to the primitive
recursive functions, but now all functions are partial, implemented
using the Part
monad, and there is an additional operation, called
μ-recursion, which performs unbounded minimization.
References #
PartRec f
means that the partial function f : ℕ → ℕ
is partially recursive.
- zero : Partrec (pure 0)
- succ : Partrec ↑Nat.succ
- left : Partrec ↑fun (n : ℕ) => (unpair n).1
- right : Partrec ↑fun (n : ℕ) => (unpair n).2
- pair {f g : ℕ →. ℕ} : Partrec f → Partrec g → Partrec fun (n : ℕ) => Nat.pair <$> f n <*> g n
- comp {f g : ℕ →. ℕ} : Partrec f → Partrec g → Partrec fun (n : ℕ) => g n >>= f
- prec {f g : ℕ →. ℕ} : Partrec f → Partrec g → Partrec (unpaired fun (a n : ℕ) => Nat.rec (f a) (fun (y : ℕ) (IH : Part ℕ) => do let i ← IH g (Nat.pair a (Nat.pair y i))) n)
- rfind {f : ℕ →. ℕ} : Partrec f → Partrec fun (a : ℕ) => Nat.rfind fun (n : ℕ) => (fun (m : ℕ) => decide (m = 0)) <$> f (Nat.pair a n)
Instances For
Partially recursive partial functions α → σ
between Primcodable
types
Equations
- Partrec f = Nat.Partrec fun (n : ℕ) => (↑(Encodable.decode n)).bind fun (a : α) => Part.map Encodable.encode (f a)
Instances For
Partially recursive partial functions α → β → σ
between Primcodable
types
Instances For
Computable functions α → σ
between Primcodable
types:
a function is computable if and only if it is partially recursive (as a partial function)
Equations
- Computable f = Partrec ↑f
Instances For
Computable functions α → β → σ
between Primcodable
types
Equations
- Computable₂ f = Computable fun (p : α × β) => f p.1 p.2
Instances For
Alias of Computable.sumInl
.
Alias of Computable.sumInr
.
Alias of Computable.sumCasesOn
.
Alias of Partrec.optionCasesOn_right
.
Alias of Partrec.sumCasesOn_left
.
Alias of Partrec.sumCasesOn_right
.