Modelling partial recursive functions using Turing machines
This file defines a simplified basis for partial recursive functions, and a turing.TM2
model
Turing machine for evaluating these functions. This amounts to a constructive proof that every
partrec
function can be evaluated by a Turing machine.
Main definitions
A simplified basis for partrec
This section constructs the type code
, which is a data type of programs with list ℕ
input and
output, with enough expressivity to write any partial recursive function. The primitives are:
zero'
appends a0
to the input. That is,zero' v = 0 :: v
.succ
returns the successor of the head of the input, defaulting to zero if there is no head:succ [] = [1]
succ (n :: v) = [n + 1]
tail
returns the tail of the inputtail [] = []
tail (n :: v) = v
cons f fs
callsf
andfs
on the input and conses the results:cons f fs v = (f v).head :: fs v
comp f g
callsf
on the output ofg
:comp f g v = f (g v)
case f g
cases on the head of the input, callingf
org
depending on whether it is zero or a successor (similar tonat.cases_on
).case f g [] = f []
case f g (0 :: v) = f v
case f g (n+1 :: v) = g (n :: v)
fix f
callsf
repeatedly, using the head of the result off
to decide whether to callf
again or finish:fix f v = []
iff v = []
fix f v = w
iff v = 0 :: w
fix f v = fix f w
iff v = n+1 :: w
(the exact value ofn
is discarded)
This basis is convenient because it is closer to the Turing machine model - the key operations are
splitting and merging of lists of unknown length, while the messy n
-ary composition operation
from the traditional basis for partial recursive functions is absent - but it retains a
compositional semantics. The first step in transitioning to Turing machines is to make a sequential
evaluator for this basis, which we take up in the next section.
- zero' : turing.to_partrec.code
- succ : turing.to_partrec.code
- tail : turing.to_partrec.code
- cons : turing.to_partrec.code → turing.to_partrec.code → turing.to_partrec.code
- comp : turing.to_partrec.code → turing.to_partrec.code → turing.to_partrec.code
- case : turing.to_partrec.code → turing.to_partrec.code → turing.to_partrec.code
- fix : turing.to_partrec.code → turing.to_partrec.code
The type of codes for primitive recursive functions. Unlike nat.partrec.code
, this uses a set
of operations on list ℕ
. See code.eval
for a description of the behavior of the primitives.
The semantics of the code
primitives, as partial functions list ℕ →. list ℕ
. By convention
we functions that return a single result return a singleton [n]
, or in some cases n :: v
where
v
will be ignored by a subsequent function.
zero'
appends a0
to the input. That is,zero' v = 0 :: v
.succ
returns the successor of the head of the input, defaulting to zero if there is no head:succ [] = [1]
succ (n :: v) = [n + 1]
tail
returns the tail of the inputtail [] = []
tail (n :: v) = v
cons f fs
callsf
andfs
on the input and conses the results:cons f fs v = (f v).head :: fs v
comp f g
callsf
on the output ofg
:comp f g v = f (g v)
case f g
cases on the head of the input, callingf
org
depending on whether it is zero or a successor (similar tonat.cases_on
).case f g [] = f []
case f g (0 :: v) = f v
case f g (n+1 :: v) = g (n :: v)
fix f
callsf
repeatedly, using the head of the result off
to decide whether to callf
again or finish:fix f v = []
iff v = []
fix f v = w
iff v = 0 :: w
fix f v = fix f w
iff v = n+1 :: w
(the exact value ofn
is discarded)
Equations
- f.fix.eval = pfun.fix (λ (v : list ℕ), roption.map (λ (v : list ℕ), ite (v.head = 0) (sum.inl v.tail) (sum.inr v.tail)) (f.eval v))
- (f.case g).eval = λ (v : list ℕ), nat.elim (f.eval v.tail) (λ (y : ℕ) (_x : roption (list ℕ)), g.eval (y :: v.tail)) v.head
- (f.comp g).eval = λ (v : list ℕ), g.eval v >>= f.eval
- (f.cons fs).eval = λ (v : list ℕ), f.eval v >>= λ (n : list ℕ), fs.eval v >>= λ (ns : list ℕ), has_pure.pure (n.head :: ns)
- turing.to_partrec.code.tail.eval = λ (v : list ℕ), has_pure.pure v.tail
- turing.to_partrec.code.succ.eval = λ (v : list ℕ), has_pure.pure [v.head.succ]
- turing.to_partrec.code.zero'.eval = λ (v : list ℕ), has_pure.pure (0 :: v)
nil
is the constant nil function: nil v = []
.
head
gets the head of the input list: head [] = [0]
, head (n :: v) = [n]
.
zero
is the constant zero function: zero v = [0]
.
pred
returns the predecessor of the head of the input:
pred [] = [0]
, pred (0 :: v) = [0]
, pred (n+1 :: v) = [n]
.
rfind f
performs the function of the rfind
primitive of partial recursive functions.
rfind f v
returns the smallest n
such that (f (n :: v)).head = 0
.
It is implemented as:
rfind f v = pred (fix (λ (n::v), f (n::v) :: n+1 :: v) (0 :: v))
The idea is that the initial state is 0 :: v
, and the fix
keeps n :: v
as its internal state;
it calls f (n :: v)
as the exit test and n+1 :: v
as the next state. At the end we get
n+1 :: v
where n
is the desired output, and pred (n+1 :: v) = [n]
returns the result.
prec f g
implements the prec
(primitive recursion) operation of partial recursive
functions. prec f g
evaluates as:
prec f g [] = [f []]
prec f g (0 :: v) = [f v]
prec f g (n+1 :: v) = [g (n :: prec f g (n :: v) :: v)]
It is implemented as:
G (a :: b :: IH :: v) = (b :: a+1 :: b-1 :: g (a :: IH :: v) :: v)
F (0 :: f_v :: v) = (f_v :: v)
F (n+1 :: f_v :: v) = (fix G (0 :: n :: f_v :: v)).tail.tail
prec f g (a :: v) = [(F (a :: f v :: v)).head]
Because fix
always evaluates its body at least once, we must special case the 0
case to avoid
calling g
more times than necessary (which could be bad if g
diverges). If the input is
0 :: v
, then F (0 :: f v :: v) = (f v :: v)
so we return [f v]
. If the input is n+1 :: v
,
we evaluate the function from the bottom up, with initial state 0 :: n :: f v :: v
. The first
number counts up, providing arguments for the applications to g
, while the second number counts
down, providing the exit condition (this is the initial b
in the return value of G
, which is
stripped by fix
). After the fix
is complete, the final state is n :: 0 :: res :: v
where
res
is the desired result, and the rest reduces this to [res]
.
Equations
- f.prec g = let G : turing.to_partrec.code := turing.to_partrec.code.tail.cons (turing.to_partrec.code.succ.cons ((turing.to_partrec.code.pred.comp turing.to_partrec.code.tail).cons ((g.comp (turing.to_partrec.code.id.cons (turing.to_partrec.code.tail.comp turing.to_partrec.code.tail))).cons (turing.to_partrec.code.tail.comp (turing.to_partrec.code.tail.comp turing.to_partrec.code.tail))))), F : turing.to_partrec.code := turing.to_partrec.code.id.case (((turing.to_partrec.code.tail.comp turing.to_partrec.code.tail).comp G.fix).comp turing.to_partrec.code.zero') in (F.comp (turing.to_partrec.code.head.cons ((f.comp turing.to_partrec.code.tail).cons turing.to_partrec.code.tail))).cons turing.to_partrec.code.nil
From compositional semantics to sequential semantics
Our initial sequential model is designed to be as similar as possible to the compositional
semantics in terms of its primitives, but it is a sequential semantics, meaning that rather than
defining an eval c : list ℕ →. list ℕ
function for each program, defined by recursion on
programs, we have a type cfg
with a step function step : cfg → option cfg
that provides a
deterministic evaluation order. In order to do this, we introduce the notion of a continuation,
which can be viewed as a code
with a hole in it where evaluation is currently taking place.
Continuations can be assigned a list ℕ →. list ℕ
semantics as well, with the interpretation
being that given a list ℕ
result returned from the code in the hole, the remainder of the
program will evaluate to a list ℕ
final value.
The continuations are:
halt
: the empty continuation: the hole is the whole program, whatever is returned is the final result. In our notation this is just_
.cons₁ fs v k
: evaluating the first part of acons
, that isk (_ :: fs v)
, wherek
is the outer continuation.cons₂ ns k
: evaluating the second part of acons
:k (ns.head :: _)
. (Technically we don't need to hold on to all ofns
here since we are already committed to taking the head, but this is more regular.)comp f k
: evaluating the first part of a composition:k (f _)
.fix f k
: waiting for the result off
in afix f
expression:k (if _.head = 0 then _.tail else fix f (_.tail))
The type cfg
of evaluation states is:
ret k v
: we have received a result, and are now evaluating the continuationk
with resultv
; that is,k v
wherek
is ready to evaluate.halt v
: we are done and the result isv
.
The main theorem of this section is that for each code c
, the state step_normal c halt v
steps
to v'
in finitely many steps if and only if code.eval c v = some v'
.
- halt : turing.to_partrec.cont
- cons₁ : turing.to_partrec.code → list ℕ → turing.to_partrec.cont → turing.to_partrec.cont
- cons₂ : list ℕ → turing.to_partrec.cont → turing.to_partrec.cont
- comp : turing.to_partrec.code → turing.to_partrec.cont → turing.to_partrec.cont
- fix : turing.to_partrec.code → turing.to_partrec.cont → turing.to_partrec.cont
The type of continuations, built up during evaluation of a code
expression.
The semantics of a continuation.
Equations
- (turing.to_partrec.cont.fix f k).eval = λ (v : list ℕ), ite (v.head = 0) (k.eval v.tail) (f.fix.eval v.tail >>= k.eval)
- (turing.to_partrec.cont.comp f k).eval = λ (v : list ℕ), f.eval v >>= k.eval
- (turing.to_partrec.cont.cons₂ ns k).eval = λ (v : list ℕ), k.eval (ns.head :: v)
- (turing.to_partrec.cont.cons₁ fs as k).eval = λ (v : list ℕ), fs.eval as >>= λ (ns : list ℕ), k.eval (v.head :: ns)
- turing.to_partrec.cont.halt.eval = has_pure.pure
- halt : list ℕ → turing.to_partrec.cfg
- ret : turing.to_partrec.cont → list ℕ → turing.to_partrec.cfg
The semantics of a continuation.
Evaluating c : code
in a continuation k : cont
and input v : list ℕ
. This goes by
recursion on c
, building an augmented continuation and a value to pass to it.
zero' v = 0 :: v
evaluates immediately, so we return it to the parent continuationsucc v = [v.head.succ]
evaluates immediately, so we return it to the parent continuationtail v = v.tail
evaluates immediately, so we return it to the parent continuationcons f fs v = (f v).head :: fs v
requires two sub-evaluations, so we evaluatef v
in the continuationk (_.head :: fs v)
(calledcont.cons₁ fs v k
)comp f g v = f (g v)
requires two sub-evaluations, so we evaluateg v
in the continuationk (f _)
(calledcont.comp f k
)case f g v = v.head.cases_on (f v.tail) (λ n, g (n :: v.tail))
has the information needed to evaluate the case statement, so we do that and transition to eitherf v
org (n :: v.tail)
.fix f v = let v' := f v in if v'.head = 0 then k v'.tail else fix f v'.tail
needs to first evaluatef v
, so we do that and leave the rest for the continuation (calledcont.fix f k
)
Equations
- turing.to_partrec.step_normal f.fix k v = turing.to_partrec.step_normal f (turing.to_partrec.cont.fix f k) v
- turing.to_partrec.step_normal (f.case g) k v = nat.elim (turing.to_partrec.step_normal f k v.tail) (λ (y : ℕ) (_x : turing.to_partrec.cfg), turing.to_partrec.step_normal g k (y :: v.tail)) v.head
- turing.to_partrec.step_normal (f.comp g) k v = turing.to_partrec.step_normal g (turing.to_partrec.cont.comp f k) v
- turing.to_partrec.step_normal (f.cons fs) k v = turing.to_partrec.step_normal f (turing.to_partrec.cont.cons₁ fs v k) v
- turing.to_partrec.step_normal turing.to_partrec.code.tail k v = turing.to_partrec.cfg.ret k v.tail
- turing.to_partrec.step_normal turing.to_partrec.code.succ k v = turing.to_partrec.cfg.ret k [v.head.succ]
- turing.to_partrec.step_normal turing.to_partrec.code.zero' k v = turing.to_partrec.cfg.ret k (0 :: v)
Evaluating a continuation k : cont
on input v : list ℕ
. This is the second part of
evaluation, when we receive results from continuations built by step_normal
.
cont.halt v = v
, so we are done and transition to thecfg.halt v
statecont.cons₁ fs as k v = k (v.head :: fs as)
, so we evaluatefs as
now with the continuationk (v.head :: _)
(calledcons₂ v k
).cont.cons₂ ns k v = k (ns.head :: v)
, where we now have everything we need to evaluatens.head :: v
, so we return it tok
.cont.comp f k v = k (f v)
, so we callf v
withk
as the continuation.cont.fix f k v = k (if v.head = 0 then k v.tail else fix f v.tail)
, wherev
is a value, so we evaluate the if statement and either callk
withv.tail
, or callfix f v
withk
as the continuation (which immediately callsf
withcont.fix f k
as the continuation).
Equations
- turing.to_partrec.step_ret (turing.to_partrec.cont.fix f k) v = ite (v.head = 0) (turing.to_partrec.step_ret k v.tail) (turing.to_partrec.step_normal f (turing.to_partrec.cont.fix f k) v.tail)
- turing.to_partrec.step_ret (turing.to_partrec.cont.comp f k) v = turing.to_partrec.step_normal f k v
- turing.to_partrec.step_ret (turing.to_partrec.cont.cons₂ ns k) v = turing.to_partrec.step_ret k (ns.head :: v)
- turing.to_partrec.step_ret (turing.to_partrec.cont.cons₁ fs as k) v = turing.to_partrec.step_normal fs (turing.to_partrec.cont.cons₂ v k) as
- turing.to_partrec.step_ret turing.to_partrec.cont.halt v = turing.to_partrec.cfg.halt v
If we are not done (in cfg.halt
state), then we must be still stuck on a continuation, so
this main loop calls step_ret
with the new continuation. The overall step
function transitions
from one cfg
to another, only halting at the cfg.halt
state.
In order to extract a compositional semantics from the sequential execution behavior of
configurations, we observe that continuations have a monoid structure, with cont.halt
as the unit
and cont.then
as the multiplication. cont.then k₁ k₂
runs k₁
until it halts, and then takes
the result of k₁
and passes it to k₂
.
We will not prove it is associative (although it is), but we are instead interested in the
associativity law k₂ (eval c k₁) = eval c (k₁.then k₂)
. This holds at both the sequential and
compositional levels, and allows us to express running a machine without the ambient continuation
and relate it to the original machine's evaluation steps. In the literature this is usually
where one uses Turing machines embedded inside other Turing machines, but this approach allows us
to avoid changing the ambient type cfg
in the middle of the recursion.
Equations
- (turing.to_partrec.cont.fix f k).then k' = turing.to_partrec.cont.fix f (k.then k')
- (turing.to_partrec.cont.comp f k).then k' = turing.to_partrec.cont.comp f (k.then k')
- (turing.to_partrec.cont.cons₂ ns k).then k' = turing.to_partrec.cont.cons₂ ns (k.then k')
- (turing.to_partrec.cont.cons₁ fs as k).then k' = turing.to_partrec.cont.cons₁ fs as (k.then k')
- turing.to_partrec.cont.halt.then k' = k'
The then k
function is a "configuration homomorphism". Its operation on states is to append
k
to the continuation of a cfg.ret
state, and to run k
on v
if we are in the cfg.halt v
state.
Equations
- (turing.to_partrec.cfg.ret k v).then k' = turing.to_partrec.cfg.ret (k.then k') v
- (turing.to_partrec.cfg.halt v).then k' = turing.to_partrec.step_ret k' v
The step_normal
function respects the then k'
homomorphism. Note that this is an exact
equality, not a simulation; the original and embedded machines move in lock-step until the
embedded machine reaches the halt state.
The step_ret
function respects the then k'
homomorphism. Note that this is an exact
equality, not a simulation; the original and embedded machines move in lock-step until the
embedded machine reaches the halt state.
This is a temporary definition, because we will prove in code_is_ok
that it always holds.
It asserts that c
is semantically correct; that is, for any k
and v
,
eval (step_normal c k v) = eval (cfg.ret k (code.eval c v))
, as an equality of partial values
(so one diverges iff the other does).
In particular, we can let k = cont.halt
, and then this asserts that step_normal c cont.halt v
evaluates to cfg.halt (code.eval c v)
.
Equations
- c.ok = ∀ (k : turing.to_partrec.cont) (v : list ℕ), turing.eval turing.to_partrec.step (turing.to_partrec.step_normal c k v) = c.eval v >>= λ (v : list ℕ), turing.eval turing.to_partrec.step (turing.to_partrec.cfg.ret k v)
Simulating sequentialized partial recursive functions in TM2
At this point we have a sequential model of partial recursive functions: the cfg
type and
step : cfg → option cfg
function from the previous section. The key feature of this model is that
it does a finite amount of computation (in fact, an amount which is statically bounded by the size
of the program) between each step, and no individual step can diverge (unlike the compositional
semantics, where every sub-part of the computation is potentially divergent). So we can utilize the
same techniques as in the other TM simulations in computability.turing_machine
to prove that
each step corresponds to a finite number of steps in a lower level model. (We don't prove it here,
but in anticipation of the complexity class P, the simulation is actually polynomial-time as well.)
The target model is turing.TM2
, which has a fixed finite set of stacks, a bit of local storage,
with programs selected from a potentially infinite (but finitely accessible) set of program
positions, or labels Λ
, each of which executes a finite sequence of basic stack commands.
For this program we will need four stacks, each on an alphabet Γ'
like so:
inductive Γ' | Cons | cons | bit0 | bit1
We represent a number as a bit sequence, lists of numbers by putting cons
after each element, and
lists of lists of natural numbers by putting Cons
after each list. For example:
0 ~> []
1 ~> [bit1]
6 ~> [bit0, bit1, bit1]
[1, 2] ~> [bit1, cons, bit0, bit1, cons]
[[], [1, 2]] ~> [Cons, bit1, cons, bit0, bit1, cons, Cons]
The four stacks are main
, rev
, aux
, stack
. In normal mode, main
contains the input to the
current program (a list ℕ
) and stack
contains data (a list (list ℕ)
) associated to the
current continuation, and in ret
mode main
contains the value that is being passed to the
continuation and stack
contains the data for the continuation. The rev
and aux
stacks are
usually empty; rev
is used to store reversed data when e.g. moving a value from one stack to
another, while aux
is used as a temporary for a main
/stack
swap that happens during cons₁
evaluation.
The only local store we need is option Γ'
, which stores the result of the last pop
operation. (Most of our working data are natural numbers, which are too large to fit in the local
store.)
The continuations from the previous section are data-carrying, containing all the values that have
been computed and are awaiting other arguments. In order to have only a finite number of
continuations appear in the program so that they can be used in machine states, we separate the
data part (anything with type list ℕ
) from the cont
type, producing a cont'
type that lacks
this information. The data is kept on the stack
stack.
Because we want to have subroutines for e.g. moving an entire stack to another place, we use an
infinite inductive type Λ'
so that we can execute a program and then return to do something else
without having to define too many different kinds of intermediate states. (We must nevertheless
prove that only finitely many labels are accessible.) The labels are:
move p k₁ k₂ q
: move elements from stackk₁
tok₂
whilep
holds of the value being moved. The last element, that failsp
, is placed in neither stack but left in the local store. At the end of the operation,k₂
will have the elements ofk₁
in reverse order. Then doq
.clear p k q
: delete elements from stackk
untilp
is true. Likemove
, the last element is left in the local storage. Then doq
.copy q
: Move all elements fromrev
to bothmain
andstack
(in reverse order), then doq
. That is, it takes(a, b, c, d)
to(b.reverse ++ a, [], c, b.reverse ++ d)
.push k f q
: pushf s
, wheres
is the local store, to stackk
, then doq
. This is a duplicate of thepush
instruction that is part of the TM2 model, but by having a subroutine just for this purpose we can build up programs to execute inside agoto
statement, where we have the flexibility to be general recursive.read (f : option Γ' → Λ')
: go to statef s
wheres
is the local store. Again this is only here for convenience.succ q
: perform a successor operation. Assuming[n]
is encoded onmain
before,[n+1]
will be on main after. This implements successor for binary natural numbers.pred q₁ q₂
: perform a predecessor operation orcase
statement. If[]
is encoded onmain
before, then we transition toq₁
with[]
on main; if(0 :: v)
is onmain
before thenv
will be onmain
after and we transition toq₁
; and if(n+1 :: v)
is onmain
before thenn :: v
will be onmain
after and we transition toq₂
.ret k
: call continuationk
. Each continuation has its own interpretation of the data instack
and sets up the data for the next continuation.ret (cons₁ fs k)
:v :: k_data
onstack
andns
onmain
, and the next step expectsv
onmain
andns :: k_data
onstack
. So we have to do a little dance here with six reverse-moves using theaux
stack to perform a three-point swap, each of which involves two reversals.ret (cons₂ k)
:ns :: k_data
is onstack
andv
is onmain
, and we have to putns.head :: v
onmain
andk_data
onstack
. This is done using thehead
subroutine.ret (fix f k)
: This stores no data, so we just check ifmain
starts with0
and if so, remove it and callk
, otherwiseclear
the first value and callf
.ret halt
: the stack is empty, andmain
has the output. Do nothing and halt.
In addition to these basic states, we define some additional subroutines that are used in the above:
push'
,peek'
,pop'
are special versions of the builtins that use the local store to supply inputs and outputs.unrev
: special casemove ff rev main
to move everything fromrev
back tomain
. Used as a cleanup operation in several functions.move_excl p k₁ k₂ q
: same asmove
but pushes the last value read back onto the source stack.move₂ p k₁ k₂ q
: doublemove
, so that the result comes out in the right order at the target stack. Implemented asmove_excl p k rev; move ff rev k₂
. Assumes that neitherk₁
nork₂
isrev
andrev
is initially empty.head k q
: get the first natural number from stackk
and reverse-move it torev
, then clear the rest of the list atk
and thenunrev
to reverse-move the head value tomain
. This is used withk = main
to implement regularhead
, i.e. ifv
is onmain
before then[v.head]
will be onmain
after; and also withk = stack
for thecons
operation, which hasv
onmain
andns :: k_data
onstack
, and results ink_data
onstack
andns.head :: v
onmain
.tr_normal
is the main entry point, defining states that perform a givencode
computation. It mostly just dispatches to functions written above.
The main theorem of this section is tr_eval
, which asserts that for each that for each code c
,
the state init c v
steps to halt v'
in finitely many steps if and only if
code.eval c v = some v'
.
- Cons : turing.partrec_to_TM2.Γ'
- cons : turing.partrec_to_TM2.Γ'
- bit0 : turing.partrec_to_TM2.Γ'
- bit1 : turing.partrec_to_TM2.Γ'
The alphabet for the stacks in the program. bit0
and bit1
are used to represent ℕ
values
as lists of binary digits, cons
is used to separate list ℕ
values, and Cons
is used to
separate list (list ℕ)
values. See the section documentation.
- main : turing.partrec_to_TM2.K'
- rev : turing.partrec_to_TM2.K'
- aux : turing.partrec_to_TM2.K'
- stack : turing.partrec_to_TM2.K'
The four stacks used by the program. main
is used to store the input value in tr_normal
mode and the output value in Λ'.ret
mode, while stack
is used to keep all the data for the
continuations. rev
is used to store reversed lists when transferring values between stacks, and
aux
is only used once in cons₁
. See the section documentation.
- halt : turing.partrec_to_TM2.cont'
- cons₁ : turing.to_partrec.code → turing.partrec_to_TM2.cont' → turing.partrec_to_TM2.cont'
- cons₂ : turing.partrec_to_TM2.cont' → turing.partrec_to_TM2.cont'
- comp : turing.to_partrec.code → turing.partrec_to_TM2.cont' → turing.partrec_to_TM2.cont'
- fix : turing.to_partrec.code → turing.partrec_to_TM2.cont' → turing.partrec_to_TM2.cont'
Continuations as in to_partrec.cont
but with the data removed. This is done because we want
the set of all continuations in the program to be finite (so that it can ultimately be encoded into
the finite state machine of a Turing machine), but a continuation can handle a potentially infinite
number of data values during execution.
- move : (turing.partrec_to_TM2.Γ' → bool) → turing.partrec_to_TM2.K' → turing.partrec_to_TM2.K' → turing.partrec_to_TM2.Λ' → turing.partrec_to_TM2.Λ'
- clear : (turing.partrec_to_TM2.Γ' → bool) → turing.partrec_to_TM2.K' → turing.partrec_to_TM2.Λ' → turing.partrec_to_TM2.Λ'
- copy : turing.partrec_to_TM2.Λ' → turing.partrec_to_TM2.Λ'
- push : turing.partrec_to_TM2.K' → (option turing.partrec_to_TM2.Γ' → option turing.partrec_to_TM2.Γ') → turing.partrec_to_TM2.Λ' → turing.partrec_to_TM2.Λ'
- read : (option turing.partrec_to_TM2.Γ' → turing.partrec_to_TM2.Λ') → turing.partrec_to_TM2.Λ'
- succ : turing.partrec_to_TM2.Λ' → turing.partrec_to_TM2.Λ'
- pred : turing.partrec_to_TM2.Λ' → turing.partrec_to_TM2.Λ' → turing.partrec_to_TM2.Λ'
- ret : turing.partrec_to_TM2.cont' → turing.partrec_to_TM2.Λ'
The set of program positions. We make extensive use of inductive types here to let us describe
"subroutines"; for example clear p k q
is a program that clears stack k
, then does q
where
q
is another label. In order to prevent this from resulting in an infinite number of distinct
accessible states, we are careful to be non-recursive (although loops are okay). See the section
documentation for a description of all the programs.
The type of TM2 statements used by this machine.
The type of TM2 configurations used by this machine.
A predicate that detects the end of a natural number, either Γ'.cons
or Γ'.Cons
(or
implicitly the end of the list), for use in predicate-taking functions like move
and clear
.
Pop a value from the stack and place the result in local store.
Equations
- turing.partrec_to_TM2.pop' k = turing.TM2.stmt.pop k (λ (x v : option turing.partrec_to_TM2.Γ'), v)
Peek a value from the stack and place the result in local store.
Equations
- turing.partrec_to_TM2.peek' k = turing.TM2.stmt.peek k (λ (x v : option turing.partrec_to_TM2.Γ'), v)
Push the value in the local store to the given stack.
Equations
- turing.partrec_to_TM2.push' k = turing.TM2.stmt.push k (λ (x : option turing.partrec_to_TM2.Γ'), x.iget)
Move everything from the rev
stack to the main
stack (reversed).
Move elements from k₁
to k₂
while p
holds, with the last element being left on k₁
.
Equations
- turing.partrec_to_TM2.move_excl p k₁ k₂ q = turing.partrec_to_TM2.Λ'.move p k₁ k₂ (turing.partrec_to_TM2.Λ'.push k₁ id q)
Move elements from k₁
to k₂
without reversion, by performing a double move via the rev
stack.
Equations
Assuming tr_list v
is on the front of stack k
, remove it, and push v.head
onto main
.
See the section documentation.
Equations
- turing.partrec_to_TM2.head k q = turing.partrec_to_TM2.Λ'.move turing.partrec_to_TM2.nat_end k turing.partrec_to_TM2.K'.rev (turing.partrec_to_TM2.Λ'.push turing.partrec_to_TM2.K'.rev (λ (_x : option turing.partrec_to_TM2.Γ'), option.some turing.partrec_to_TM2.Γ'.cons) (turing.partrec_to_TM2.Λ'.read (λ (s : option turing.partrec_to_TM2.Γ'), ite (s = option.some turing.partrec_to_TM2.Γ'.Cons) id (turing.partrec_to_TM2.Λ'.clear (λ (x : turing.partrec_to_TM2.Γ'), decidable.to_bool (x = turing.partrec_to_TM2.Γ'.Cons)) k) (turing.partrec_to_TM2.unrev q))))
The program that evaluates code c
with continuation k
. This expects an initial state where
tr_list v
is on main
, tr_cont_stack k
is on stack
, and aux
and rev
are empty.
See the section documentation for details.
Equations
- turing.partrec_to_TM2.tr_normal f.fix k = turing.partrec_to_TM2.tr_normal f (turing.partrec_to_TM2.cont'.fix f k)
- turing.partrec_to_TM2.tr_normal (f.case g) k = (turing.partrec_to_TM2.tr_normal f k).pred (turing.partrec_to_TM2.tr_normal g k)
- turing.partrec_to_TM2.tr_normal (f.comp g) k = turing.partrec_to_TM2.tr_normal g (turing.partrec_to_TM2.cont'.comp f k)
- turing.partrec_to_TM2.tr_normal (f.cons fs) k = turing.partrec_to_TM2.Λ'.push turing.partrec_to_TM2.K'.stack (λ (_x : option turing.partrec_to_TM2.Γ'), option.some turing.partrec_to_TM2.Γ'.Cons) (turing.partrec_to_TM2.Λ'.move (λ (_x : turing.partrec_to_TM2.Γ'), bool.ff) turing.partrec_to_TM2.K'.main turing.partrec_to_TM2.K'.rev (turing.partrec_to_TM2.tr_normal f (turing.partrec_to_TM2.cont'.cons₁ fs k)).copy)
- turing.partrec_to_TM2.tr_normal turing.to_partrec.code.tail k = turing.partrec_to_TM2.Λ'.clear turing.partrec_to_TM2.nat_end turing.partrec_to_TM2.K'.main (turing.partrec_to_TM2.Λ'.ret k)
- turing.partrec_to_TM2.tr_normal turing.to_partrec.code.succ k = turing.partrec_to_TM2.head turing.partrec_to_TM2.K'.main (turing.partrec_to_TM2.Λ'.ret k).succ
- turing.partrec_to_TM2.tr_normal turing.to_partrec.code.zero' k = turing.partrec_to_TM2.Λ'.push turing.partrec_to_TM2.K'.main (λ (_x : option turing.partrec_to_TM2.Γ'), option.some turing.partrec_to_TM2.Γ'.cons) (turing.partrec_to_TM2.Λ'.ret k)
The main program. See the section documentation for details.
Equations
- turing.partrec_to_TM2.tr (turing.partrec_to_TM2.Λ'.ret (turing.partrec_to_TM2.cont'.fix f k)) = turing.partrec_to_TM2.pop' turing.partrec_to_TM2.K'.main (turing.TM2.stmt.goto (λ (s : option turing.partrec_to_TM2.Γ'), cond (turing.partrec_to_TM2.nat_end s.iget) (turing.partrec_to_TM2.Λ'.ret k) (turing.partrec_to_TM2.Λ'.clear turing.partrec_to_TM2.nat_end turing.partrec_to_TM2.K'.main (turing.partrec_to_TM2.tr_normal f (turing.partrec_to_TM2.cont'.fix f k)))))
- turing.partrec_to_TM2.tr (turing.partrec_to_TM2.Λ'.ret (turing.partrec_to_TM2.cont'.comp f k)) = turing.TM2.stmt.goto (λ (_x : option turing.partrec_to_TM2.Γ'), turing.partrec_to_TM2.tr_normal f k)
- turing.partrec_to_TM2.tr (turing.partrec_to_TM2.Λ'.ret k.cons₂) = turing.TM2.stmt.goto (λ (_x : option turing.partrec_to_TM2.Γ'), turing.partrec_to_TM2.head turing.partrec_to_TM2.K'.stack (turing.partrec_to_TM2.Λ'.ret k))
- turing.partrec_to_TM2.tr (turing.partrec_to_TM2.Λ'.ret (turing.partrec_to_TM2.cont'.cons₁ fs k)) = turing.TM2.stmt.goto (λ (_x : option turing.partrec_to_TM2.Γ'), turing.partrec_to_TM2.move₂ (λ (_x : turing.partrec_to_TM2.Γ'), bool.ff) turing.partrec_to_TM2.K'.main turing.partrec_to_TM2.K'.aux (turing.partrec_to_TM2.move₂ (λ (s : turing.partrec_to_TM2.Γ'), decidable.to_bool (s = turing.partrec_to_TM2.Γ'.Cons)) turing.partrec_to_TM2.K'.stack turing.partrec_to_TM2.K'.main (turing.partrec_to_TM2.move₂ (λ (_x : turing.partrec_to_TM2.Γ'), bool.ff) turing.partrec_to_TM2.K'.aux turing.partrec_to_TM2.K'.stack (turing.partrec_to_TM2.tr_normal fs k.cons₂))))
- turing.partrec_to_TM2.tr (turing.partrec_to_TM2.Λ'.ret turing.partrec_to_TM2.cont'.halt) = turing.TM2.stmt.load (λ (_x : option turing.partrec_to_TM2.Γ'), option.none) turing.TM2.stmt.halt
- turing.partrec_to_TM2.tr (q₁.pred q₂) = turing.partrec_to_TM2.pop' turing.partrec_to_TM2.K'.main (turing.TM2.stmt.branch (λ (s : option turing.partrec_to_TM2.Γ'), decidable.to_bool (s = option.some turing.partrec_to_TM2.Γ'.bit0)) (turing.TM2.stmt.push turing.partrec_to_TM2.K'.rev (λ (_x : option turing.partrec_to_TM2.Γ'), turing.partrec_to_TM2.Γ'.bit1) (turing.TM2.stmt.goto (λ (_x : option turing.partrec_to_TM2.Γ'), q₁.pred q₂))) (turing.TM2.stmt.branch (λ (s : option turing.partrec_to_TM2.Γ'), turing.partrec_to_TM2.nat_end s.iget) (turing.TM2.stmt.goto (λ (_x : option turing.partrec_to_TM2.Γ'), q₁)) (turing.partrec_to_TM2.peek' turing.partrec_to_TM2.K'.main (turing.TM2.stmt.branch (λ (s : option turing.partrec_to_TM2.Γ'), turing.partrec_to_TM2.nat_end s.iget) (turing.TM2.stmt.goto (λ (_x : option turing.partrec_to_TM2.Γ'), turing.partrec_to_TM2.unrev q₂)) (turing.TM2.stmt.push turing.partrec_to_TM2.K'.rev (λ (_x : option turing.partrec_to_TM2.Γ'), turing.partrec_to_TM2.Γ'.bit0) (turing.TM2.stmt.goto (λ (_x : option turing.partrec_to_TM2.Γ'), turing.partrec_to_TM2.unrev q₂)))))))
- turing.partrec_to_TM2.tr q.succ = turing.partrec_to_TM2.pop' turing.partrec_to_TM2.K'.main (turing.TM2.stmt.branch (λ (s : option turing.partrec_to_TM2.Γ'), decidable.to_bool (s = option.some turing.partrec_to_TM2.Γ'.bit1)) (turing.TM2.stmt.push turing.partrec_to_TM2.K'.rev (λ (_x : option turing.partrec_to_TM2.Γ'), turing.partrec_to_TM2.Γ'.bit0) (turing.TM2.stmt.goto (λ (_x : option turing.partrec_to_TM2.Γ'), q.succ))) (turing.TM2.stmt.branch (λ (s : option turing.partrec_to_TM2.Γ'), decidable.to_bool (s = option.some turing.partrec_to_TM2.Γ'.cons)) (turing.TM2.stmt.push turing.partrec_to_TM2.K'.main (λ (_x : option turing.partrec_to_TM2.Γ'), turing.partrec_to_TM2.Γ'.cons) (turing.TM2.stmt.push turing.partrec_to_TM2.K'.main (λ (_x : option turing.partrec_to_TM2.Γ'), turing.partrec_to_TM2.Γ'.bit1) (turing.TM2.stmt.goto (λ (_x : option turing.partrec_to_TM2.Γ'), turing.partrec_to_TM2.unrev q)))) (turing.TM2.stmt.push turing.partrec_to_TM2.K'.main (λ (_x : option turing.partrec_to_TM2.Γ'), turing.partrec_to_TM2.Γ'.bit1) (turing.TM2.stmt.goto (λ (_x : option turing.partrec_to_TM2.Γ'), turing.partrec_to_TM2.unrev q)))))
- turing.partrec_to_TM2.tr (turing.partrec_to_TM2.Λ'.read q) = turing.TM2.stmt.goto q
- turing.partrec_to_TM2.tr (turing.partrec_to_TM2.Λ'.push k f q) = turing.TM2.stmt.branch (λ (s : option turing.partrec_to_TM2.Γ'), (f s).is_some) (turing.TM2.stmt.push k (λ (s : option turing.partrec_to_TM2.Γ'), (f s).iget) (turing.TM2.stmt.goto (λ (_x : option turing.partrec_to_TM2.Γ'), q))) (turing.TM2.stmt.goto (λ (_x : option turing.partrec_to_TM2.Γ'), q))
- turing.partrec_to_TM2.tr q.copy = turing.partrec_to_TM2.pop' turing.partrec_to_TM2.K'.rev (turing.TM2.stmt.branch option.is_some (turing.partrec_to_TM2.push' turing.partrec_to_TM2.K'.main (turing.partrec_to_TM2.push' turing.partrec_to_TM2.K'.stack (turing.TM2.stmt.goto (λ (_x : option turing.partrec_to_TM2.Γ'), q.copy)))) (turing.TM2.stmt.goto (λ (_x : option turing.partrec_to_TM2.Γ'), q)))
- turing.partrec_to_TM2.tr (turing.partrec_to_TM2.Λ'.clear p k q) = turing.partrec_to_TM2.pop' k (turing.TM2.stmt.branch (λ (s : option turing.partrec_to_TM2.Γ'), s.elim bool.tt p) (turing.TM2.stmt.goto (λ (_x : option turing.partrec_to_TM2.Γ'), q)) (turing.TM2.stmt.goto (λ (_x : option turing.partrec_to_TM2.Γ'), turing.partrec_to_TM2.Λ'.clear p k q)))
- turing.partrec_to_TM2.tr (turing.partrec_to_TM2.Λ'.move p k₁ k₂ q) = turing.partrec_to_TM2.pop' k₁ (turing.TM2.stmt.branch (λ (s : option turing.partrec_to_TM2.Γ'), s.elim bool.tt p) (turing.TM2.stmt.goto (λ (_x : option turing.partrec_to_TM2.Γ'), q)) (turing.partrec_to_TM2.push' k₂ (turing.TM2.stmt.goto (λ (_x : option turing.partrec_to_TM2.Γ'), turing.partrec_to_TM2.Λ'.move p k₁ k₂ q))))
Translating a cont
continuation to a cont'
continuation simply entails dropping all the
data. This data is instead encoded in tr_cont_stack
in the configuration.
Equations
- turing.partrec_to_TM2.tr_cont (turing.to_partrec.cont.fix c k) = turing.partrec_to_TM2.cont'.fix c (turing.partrec_to_TM2.tr_cont k)
- turing.partrec_to_TM2.tr_cont (turing.to_partrec.cont.comp c k) = turing.partrec_to_TM2.cont'.comp c (turing.partrec_to_TM2.tr_cont k)
- turing.partrec_to_TM2.tr_cont (turing.to_partrec.cont.cons₂ _x k) = (turing.partrec_to_TM2.tr_cont k).cons₂
- turing.partrec_to_TM2.tr_cont (turing.to_partrec.cont.cons₁ c _x k) = turing.partrec_to_TM2.cont'.cons₁ c (turing.partrec_to_TM2.tr_cont k)
- turing.partrec_to_TM2.tr_cont turing.to_partrec.cont.halt = turing.partrec_to_TM2.cont'.halt
We use pos_num
to define the translation of binary natural numbers. A natural number is
represented as a little-endian list of bit0
and bit1
elements:
1 = [bit1]
2 = [bit0, bit1]
3 = [bit1, bit1]
4 = [bit0, bit0, bit1]
In particular, this representation guarantees no trailing bit0
's at the end of the list.
Equations
Lists are translated with a cons
after each encoded number.
For example:
[] = []
[0] = [cons]
[1] = [bit1, cons]
[6, 0] = [bit0, bit1, bit1, cons, cons]
Lists of lists are translated with a Cons
after each encoded list.
For example:
[] = []
[[]] = [Cons]
[[], []] = [Cons, Cons]
[[0]] = [cons, Cons]
[[1, 2], [0]] = [bit1, cons, bit0, bit1, cons, Cons, cons, Cons]
The data part of a continuation is a list of lists, which is encoded on the stack
stack
using tr_llist
.
Equations
- turing.partrec_to_TM2.cont_stack (turing.to_partrec.cont.fix _x k) = turing.partrec_to_TM2.cont_stack k
- turing.partrec_to_TM2.cont_stack (turing.to_partrec.cont.comp _x k) = turing.partrec_to_TM2.cont_stack k
- turing.partrec_to_TM2.cont_stack (turing.to_partrec.cont.cons₂ ns k) = ns :: turing.partrec_to_TM2.cont_stack k
- turing.partrec_to_TM2.cont_stack (turing.to_partrec.cont.cons₁ _x ns k) = ns :: turing.partrec_to_TM2.cont_stack k
- turing.partrec_to_TM2.cont_stack turing.to_partrec.cont.halt = list.nil
The data part of a continuation is a list of lists, which is encoded on the stack
stack
using tr_llist
.
This is the nondependent eliminator for K'
, but we use it specifically here in order to
represent the stack data as four lists rather than as a function K' → list Γ'
, because this makes
rewrites easier. The theorems K'.elim_update_main
et. al. show how such a function is updated
after an update
to one of the components.
Equations
The halting state corresponding to a list ℕ
output value.
The cfg
states map to cfg'
states almost one to one, except that in normal operation the
local store contains an arbitrary garbage value. To make the final theorem cleaner we explicitly
clear it in the halt state so that there is exactly one configuration corresponding to output v
.
Equations
- turing.partrec_to_TM2.tr_cfg (turing.to_partrec.cfg.ret k v) c' = ∃ (s : option turing.partrec_to_TM2.Γ'), c' = {l := option.some (turing.partrec_to_TM2.Λ'.ret (turing.partrec_to_TM2.tr_cont k)), var := s, stk := turing.partrec_to_TM2.K'.elim (turing.partrec_to_TM2.tr_list v) list.nil list.nil (turing.partrec_to_TM2.tr_cont_stack k)}
- turing.partrec_to_TM2.tr_cfg (turing.to_partrec.cfg.halt v) c' = (c' = turing.partrec_to_TM2.halt v)
This could be a general list definition, but it is also somewhat specialized to this
application. split_at_pred p L
will search L
for the first element satisfying p
.
If it is found, say L = l₁ ++ a :: l₂
where a
satisfies p
but l₁
does not, then it returns
(l₁, some a, l₂)
. Otherwise, if there is no such element, it returns (L, none, [])
.
Equations
- turing.partrec_to_TM2.split_at_pred p (a :: as) = cond (p a) (list.nil α, option.some a, as) (turing.partrec_to_TM2.split_at_pred._match_1 a (turing.partrec_to_TM2.split_at_pred p as))
- turing.partrec_to_TM2.split_at_pred p list.nil = (list.nil α, option.none α, list.nil α)
- turing.partrec_to_TM2.split_at_pred._match_1 a (l₁, o, l₂) = (a :: l₁, o, l₂)
The initial state, evaluating function c
on input v
.