Turing machines #
The files PostTuringMachine.lean
and TuringMachine.lean
define
a sequence of simple machine languages, starting with Turing machines and working
up to more complex languages based on Wang B-machines.
PostTuringMachine.lean
covers the TM0 model and TM1 model;
TuringMachine.lean
adds the TM2 model.
Naming conventions #
Each model of computation in this file shares a naming convention for the elements of a model of computation. These are the parameters for the language:
Γ
is the alphabet on the tape.Λ
is the set of labels, or internal machine states.σ
is the type of internal memory, not on the tape. This does not exist in the TM0 model, and later models achieve this by mixing it intoΛ
.K
is used in the TM2 model, which has multiple stacks, and denotes the number of such stacks.
All of these variables denote "essentially finite" types, but for technical reasons it is convenient to allow them to be infinite anyway. When using an infinite type, we will be interested to prove that only finitely many values of the type are ever interacted with.
Given these parameters, there are a few common structures for the model that arise:
Stmt
is the set of all actions that can be performed in one step. For the TM0 model this set is finite, and for later models it is an infinite inductive type representing "possible program texts".Cfg
is the set of instantaneous configurations, that is, the state of the machine together with its environment.Machine
is the set of all machines in the model. Usually this is approximately a functionΛ → Stmt
, although different models have different ways of halting and other actions.step : Cfg → Option Cfg
is the function that describes how the state evolves over one step. Ifstep c = none
, thenc
is a terminal state, and the result of the computation is read off fromc
. Because of the type ofstep
, these models are all deterministic by construction.init : Input → Cfg
sets up the initial state. The typeInput
depends on the model; in most cases it isList Γ
.eval : Machine → Input → Part Output
, given a machineM
and inputi
, starts frominit i
, runsstep
until it reaches an output, and then applies a functionCfg → Output
to the final state to obtain the result. The typeOutput
depends on the model.Supports : Machine → Finset Λ → Prop
asserts that a machineM
starts inS : Finset Λ
, and can only ever jump to other states insideS
. This implies that the behavior ofM
on any input cannot depend on its values outsideS
. We use this to allowΛ
to be an infinite set when convenient, and prove that only finitely many of these states are actually accessible. This formalizes "essentially finite" mentioned above.
The reflexive transitive closure of a state transition function. Reaches f a b
means
there is a finite sequence of steps f a = some a₁
, f a₁ = some a₂
, ... such that aₙ = b
.
This relation permits zero steps of the state transition function.
Equations
- Turing.Reaches f = Relation.ReflTransGen fun (a b : σ) => b ∈ f a
Instances For
The transitive closure of a state transition function. Reaches₁ f a b
means there is a
nonempty finite sequence of steps f a = some a₁
, f a₁ = some a₂
, ... such that aₙ = b
.
This relation does not permit zero steps of the state transition function.
Equations
- Turing.Reaches₁ f = Relation.TransGen fun (a b : σ) => b ∈ f a
Instances For
A variation on Reaches
. Reaches₀ f a b
holds if whenever Reaches₁ f b c
then
Reaches₁ f a c
. This is a weaker property than Reaches
and is useful for replacing states with
equivalent states without taking a step.
Equations
- Turing.Reaches₀ f a b = ∀ (c : σ), Turing.Reaches₁ f b c → Turing.Reaches₁ f a c
Instances For
(co-)Induction principle for eval
. If a property C
holds of any point a
evaluating to b
which is either terminal (meaning a = b
) or where the next point also satisfies C
, then it
holds of any point where eval f a
evaluates to b
. This formalizes the notion that if
eval f a
evaluates to b
then it reaches terminal state b
in finitely many steps.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Given a relation tr : σ₁ → σ₂ → Prop
between state spaces, and state transition functions
f₁ : σ₁ → Option σ₁
and f₂ : σ₂ → Option σ₂
, Respects f₁ f₂ tr
means that if tr a₁ a₂
holds
initially and f₁
takes a step to a₂
then f₂
will take one or more steps before reaching a
state b₂
satisfying tr a₂ b₂
, and if f₁ a₁
terminates then f₂ a₂
also terminates.
Such a relation tr
is also known as a refinement.
Equations
- Turing.Respects f₁ f₂ tr = ∀ ⦃a₁ : σ₁⦄ ⦃a₂ : σ₂⦄, tr a₁ a₂ → match f₁ a₁ with | some b₁ => ∃ (b₂ : σ₂), tr b₁ b₂ ∧ Turing.Reaches₁ f₂ a₂ b₂ | none => f₂ a₂ = none
Instances For
A simpler version of Respects
when the state transition relation tr
is a function.
Equations
- Turing.FRespects f₂ tr a₂ (some b₁) = Turing.Reaches₁ f₂ a₂ (tr b₁)
- Turing.FRespects f₂ tr a₂ none = (f₂ a₂ = none)
Instances For
The TM0 model #
A TM0 Turing machine is essentially a Post-Turing machine, adapted for type theory.
A Post-Turing machine with symbol type Γ
and label type Λ
is a function
Λ → Γ → Option (Λ × Stmt)
, where a Stmt
can be either move left
, move right
or write a
for a : Γ
. The machine works over a "tape", a doubly-infinite sequence of elements of Γ
, and
an instantaneous configuration, Cfg
, is a label q : Λ
indicating the current internal state of
the machine, and a Tape Γ
(which is essentially ℤ →₀ Γ
). The evolution is described by the
step
function:
- If
M q T.head = none
, then the machine halts. - If
M q T.head = some (q', s)
, then the machine performs actions : Stmt
and then transitions to stateq'
.
The initial state takes a List Γ
and produces a Tape Γ
where the head of the list is the head
of the tape and the rest of the list extends to the right, with the left side all blank. The final
state takes the entire right side of the tape right or equal to the current position of the
machine. (This is actually a ListBlank Γ
, not a List Γ
, because we don't know, at this level
of generality, where the output ends. If equality to default : Γ
is decidable we can trim the list
to remove the infinite tail of blanks.)
Equations
- Turing.TM0.Stmt.inhabited Γ = { default := Turing.TM0.Stmt.write default }
A Post-Turing machine with symbol type Γ
and label type Λ
is a function which, given the current state q : Λ
and
the tape head a : Γ
, either halts (returns none
) or returns
a new state q' : Λ
and a Stmt
describing what to do,
either a move left or right, or a write command.
Both Λ
and Γ
are required to be inhabited; the default value
for Γ
is the "blank" tape value, and the default value of Λ
is
the initial state.
Equations
- Turing.TM0.Machine Γ Λ = (Λ → Γ → Option (Λ × Turing.TM0.Stmt Γ))
Instances For
Equations
The configuration state of a Turing machine during operation
consists of a label (machine state), and a tape.
The tape is represented in the form (a, L, R)
, meaning the tape looks like L.rev ++ [a] ++ R
with the machine currently reading the a
. The lists are
automatically extended with blanks as the machine moves around.
- q : Λ
The current machine state.
- Tape : Turing.Tape Γ
The current state of the tape: current symbol, left and right parts.
Instances For
The statement Reaches M s₁ s₂
means that s₂
is obtained
starting from s₁
after a finite number of steps from s₂
.
Equations
- Turing.TM0.Reaches M = Relation.ReflTransGen fun (a b : Turing.TM0.Cfg Γ Λ) => b ∈ Turing.TM0.step M a
Instances For
The initial configuration.
Equations
- Turing.TM0.init l = { q := default, Tape := Turing.Tape.mk₁ l }
Instances For
Evaluate a Turing machine on initial input to a final state, if it terminates.
Equations
- Turing.TM0.eval M l = Part.map (fun (c : Turing.TM0.Cfg Γ Λ) => c.Tape.right₀) (Turing.eval (Turing.TM0.step M) (Turing.TM0.init l))
Instances For
The raw definition of a Turing machine does not require that
Γ
and Λ
are finite, and in practice we will be interested
in the infinite Λ
case. We recover instead a notion of
"effectively finite" Turing machines, which only make use of a
finite subset of their states. We say that a set S ⊆ Λ
supports a Turing machine M
if S
is closed under the
transition function and contains the initial state.
Equations
Instances For
Map a TM statement across a function. This does nothing to move statements and maps the write values.
Equations
- Turing.TM0.Stmt.map f (Turing.TM0.Stmt.move d) = Turing.TM0.Stmt.move d
- Turing.TM0.Stmt.map f (Turing.TM0.Stmt.write a_1) = Turing.TM0.Stmt.write (f.f a_1)
Instances For
Map a configuration across a function, given f : Γ → Γ'
a map of the alphabets and
g : Λ → Λ'
a map of the machine states.
Equations
- Turing.TM0.Cfg.map f g { q := q, Tape := T } = { q := g q, Tape := Turing.Tape.map f T }
Instances For
Because the state transition function uses the alphabet and machine states in both the input
and output, to map a machine from one alphabet and machine state space to another we need functions
in both directions, essentially an Equiv
without the laws.
Equations
- M.map f₁ f₂ g₁ g₂ x✝¹ x✝ = Option.map (Prod.map g₁ (Turing.TM0.Stmt.map f₁)) (M (g₂ x✝¹) (f₂.f x✝))
Instances For
The TM1 model #
The TM1 model is a simplification and extension of TM0 (Post-Turing model) in the direction of
Wang B-machines. The machine's internal state is extended with a (finite) store σ
of variables
that may be accessed and updated at any time.
A machine is given by a Λ
indexed set of procedures or functions. Each function has a body which
is a Stmt
. Most of the regular commands are allowed to use the current value a
of the local
variables and the value T.head
on the tape to calculate what to write or how to change local
state, but the statements themselves have a fixed structure. The Stmt
s can be as follows:
move d q
: move left or right, and then doq
write (f : Γ → σ → Γ) q
: writef a T.head
to the tape, then doq
load (f : Γ → σ → σ) q
: change the internal state tof a T.head
branch (f : Γ → σ → Bool) qtrue qfalse
: Iff a T.head
is true, doqtrue
, elseqfalse
goto (f : Γ → σ → Λ)
: Go to labelf a T.head
halt
: Transition to the halting state, which halts on the following step
Note that here most statements do not have labels; goto
commands can only go to a new function.
Only the goto
and halt
statements actually take a step; the rest is done by recursion on
statements and so take 0 steps. (There is a uniform bound on how many statements can be executed
before the next goto
, so this is an O(1)
speedup with the constant depending on the machine.)
The halt
command has a one step stutter before actually halting so that any changes made before
the halt have a chance to be "committed", since the eval
relation uses the final configuration
before the halt as the output, and move
and write
etc. take 0 steps in this model.
The TM1 model is a simplification and extension of TM0
(Post-Turing model) in the direction of Wang B-machines. The machine's
internal state is extended with a (finite) store σ
of variables
that may be accessed and updated at any time.
A machine is given by a Λ
indexed set of procedures or functions.
Each function has a body which is a Stmt
, which can either be a
move
or write
command, a branch
(if statement based on the
current tape value), a load
(set the variable value),
a goto
(call another function), or halt
. Note that here
most statements do not have labels; goto
commands can only
go to a new function. All commands have access to the variable value
and current tape value.
- move {Γ : Type u_1} {Λ : Type u_2} {σ : Type u_3} : Dir → Stmt Γ Λ σ → Stmt Γ Λ σ
- write {Γ : Type u_1} {Λ : Type u_2} {σ : Type u_3} : (Γ → σ → Γ) → Stmt Γ Λ σ → Stmt Γ Λ σ
- load {Γ : Type u_1} {Λ : Type u_2} {σ : Type u_3} : (Γ → σ → σ) → Stmt Γ Λ σ → Stmt Γ Λ σ
- branch {Γ : Type u_1} {Λ : Type u_2} {σ : Type u_3} : (Γ → σ → Bool) → Stmt Γ Λ σ → Stmt Γ Λ σ → Stmt Γ Λ σ
- goto {Γ : Type u_1} {Λ : Type u_2} {σ : Type u_3} : (Γ → σ → Λ) → Stmt Γ Λ σ
- halt {Γ : Type u_1} {Λ : Type u_2} {σ : Type u_3} : Stmt Γ Λ σ
Instances For
Equations
- Turing.TM1.Stmt.inhabited Γ Λ σ = { default := Turing.TM1.Stmt.halt }
The configuration of a TM1 machine is given by the currently evaluating statement, the variable store value, and the tape.
- l : Option Λ
The statement (if any) which is currently evaluated
- var : σ
The current value of the variable store
- Tape : Turing.Tape Γ
The current state of the tape
Instances For
The semantics of TM1 evaluation.
Equations
- Turing.TM1.stepAux (Turing.TM1.Stmt.move d q) x✝¹ x✝ = Turing.TM1.stepAux q x✝¹ (Turing.Tape.move d x✝)
- Turing.TM1.stepAux (Turing.TM1.Stmt.write a q) x✝¹ x✝ = Turing.TM1.stepAux q x✝¹ (Turing.Tape.write (a x✝.head x✝¹) x✝)
- Turing.TM1.stepAux (Turing.TM1.Stmt.load s q) x✝¹ x✝ = Turing.TM1.stepAux q (s x✝.head x✝¹) x✝
- Turing.TM1.stepAux (Turing.TM1.Stmt.branch p q₁ q₂) x✝¹ x✝ = bif p x✝.head x✝¹ then Turing.TM1.stepAux q₁ x✝¹ x✝ else Turing.TM1.stepAux q₂ x✝¹ x✝
- Turing.TM1.stepAux (Turing.TM1.Stmt.goto l) x✝¹ x✝ = { l := some (l x✝.head x✝¹), var := x✝¹, Tape := x✝ }
- Turing.TM1.stepAux Turing.TM1.Stmt.halt x✝¹ x✝ = { l := none, var := x✝¹, Tape := x✝ }
Instances For
The state transition function.
Equations
- Turing.TM1.step M { l := none, var := var, Tape := Tape } = none
- Turing.TM1.step M { l := some l, var := v, Tape := T } = some (Turing.TM1.stepAux (M l) v T)
Instances For
A set S
of labels supports the statement q
if all the goto
statements in q
refer only to other functions in S
.
Equations
- Turing.TM1.SupportsStmt S (Turing.TM1.Stmt.move a q) = Turing.TM1.SupportsStmt S q
- Turing.TM1.SupportsStmt S (Turing.TM1.Stmt.write a q) = Turing.TM1.SupportsStmt S q
- Turing.TM1.SupportsStmt S (Turing.TM1.Stmt.load a q) = Turing.TM1.SupportsStmt S q
- Turing.TM1.SupportsStmt S (Turing.TM1.Stmt.branch a q₁ q₂) = (Turing.TM1.SupportsStmt S q₁ ∧ Turing.TM1.SupportsStmt S q₂)
- Turing.TM1.SupportsStmt S (Turing.TM1.Stmt.goto l) = ∀ (a : Γ) (v : σ), l a v ∈ S
- Turing.TM1.SupportsStmt S Turing.TM1.Stmt.halt = True
Instances For
The subterm closure of a statement.
Equations
- Turing.TM1.stmts₁ (Turing.TM1.Stmt.move a q) = insert (Turing.TM1.Stmt.move a q) (Turing.TM1.stmts₁ q)
- Turing.TM1.stmts₁ (Turing.TM1.Stmt.write a q) = insert (Turing.TM1.Stmt.write a q) (Turing.TM1.stmts₁ q)
- Turing.TM1.stmts₁ (Turing.TM1.Stmt.load a q) = insert (Turing.TM1.Stmt.load a q) (Turing.TM1.stmts₁ q)
- Turing.TM1.stmts₁ (Turing.TM1.Stmt.branch a q₁ q₂) = insert (Turing.TM1.Stmt.branch a q₁ q₂) (Turing.TM1.stmts₁ q₁ ∪ Turing.TM1.stmts₁ q₂)
- Turing.TM1.stmts₁ x✝ = {x✝}
Instances For
The set of all statements in a Turing machine, plus one extra value none
representing the
halt state. This is used in the TM1 to TM0 reduction.
Equations
- Turing.TM1.stmts M S = Finset.insertNone (S.biUnion fun (q : Λ) => Turing.TM1.stmts₁ (M q))
Instances For
A set S
of labels supports machine M
if all the goto
statements in the functions in S
refer only to other functions
in S
.
Equations
- Turing.TM1.Supports M S = (default ∈ S ∧ ∀ q ∈ S, Turing.TM1.SupportsStmt S (M q))
Instances For
The initial state, given a finite input that is placed on the tape starting at the TM head and going to the right.
Equations
- Turing.TM1.init l = { l := some default, var := default, Tape := Turing.Tape.mk₁ l }
Instances For
Evaluate a TM to completion, resulting in an output list on the tape (with an indeterminate number of blanks on the end).
Equations
- Turing.TM1.eval M l = Part.map (fun (c : Turing.TM1.Cfg Γ Λ σ) => c.Tape.right₀) (Turing.eval (Turing.TM1.step M) (Turing.TM1.init l))
Instances For
TM1 emulator in TM0 #
To prove that TM1 computable functions are TM0 computable, we need to reduce each TM1 program to a TM0 program. So suppose a TM1 program is given. We take the following:
- The alphabet
Γ
is the same for both TM1 and TM0 - The set of states
Λ'
is defined to beOption Stmt₁ × σ
, that is, a TM1 statement ornone
representing halt, and the possible settings of the internal variables. Note that this is an infinite set, becauseStmt₁
is infinite. This is okay because we assume that from the initial TM1 state, only finitely many other labels are reachable, and there are only finitely many statements that appear in all of these functions.
Even though Stmt₁
contains a statement called halt
, we must separate it from none
(some halt
steps to none
and none
actually halts) because there is a one step stutter in the
TM1 semantics.
The base machine state space is a pair of an Option Stmt₁
representing the current program
to be executed, or none
for the halt state, and a σ
which is the local state (stored in the TM,
not the tape). Because there are an infinite number of programs, this state space is infinite, but
for a finitely supported TM1 machine and a finite type σ
, only finitely many of these states are
reachable.
Equations
- Turing.TM1to0.Λ' M = (Option (Turing.TM1.Stmt Γ Λ σ) × σ)
Instances For
The core TM1 → TM0 translation function. Here s
is the current value on the tape, and the
Stmt₁
is the TM1 statement to translate, with local state v : σ
. We evaluate all regular
instructions recursively until we reach either a move
or write
command, or a goto
; in the
latter case we emit a dummy write s
step and transition to the new target location.
Equations
- Turing.TM1to0.trAux M s (Turing.TM1.Stmt.move d q) x✝ = ((some q, x✝), Turing.TM0.Stmt.move d)
- Turing.TM1to0.trAux M s (Turing.TM1.Stmt.write a q) x✝ = ((some q, x✝), Turing.TM0.Stmt.write (a s x✝))
- Turing.TM1to0.trAux M s (Turing.TM1.Stmt.load a q) x✝ = Turing.TM1to0.trAux M s q (a s x✝)
- Turing.TM1to0.trAux M s (Turing.TM1.Stmt.branch p q₁ q₂) x✝ = bif p s x✝ then Turing.TM1to0.trAux M s q₁ x✝ else Turing.TM1to0.trAux M s q₂ x✝
- Turing.TM1to0.trAux M s (Turing.TM1.Stmt.goto l) x✝ = ((some (M (l s x✝)), x✝), Turing.TM0.Stmt.write s)
- Turing.TM1to0.trAux M s Turing.TM1.Stmt.halt x✝ = ((none, x✝), Turing.TM0.Stmt.write s)
Instances For
The translated TM0 machine (given the TM1 machine input).
Equations
- Turing.TM1to0.tr M (none, snd) x✝ = none
- Turing.TM1to0.tr M (some q, v) x✝ = some (Turing.TM1to0.trAux M x✝ q v)
Instances For
Translate configurations from TM1 to TM0.
Equations
- Turing.TM1to0.trCfg M { l := l, var := v, Tape := T } = { q := (Option.map M l, v), Tape := T }
Instances For
Given a finite set of accessible Λ
machine states, there is a finite set of accessible
machine states in the target (even though the type Λ'
is infinite).
Equations
Instances For
TM1(Γ) emulator in TM1(Bool) #
The most parsimonious Turing machine model that is still Turing complete is TM0
with Γ = Bool
.
Because our construction in the previous section reducing TM1
to TM0
doesn't change the
alphabet, we can do the alphabet reduction on TM1
instead of TM0
directly.
The basic idea is to use a bijection between Γ
and a subset of Vector Bool n
, where n
is a
fixed constant. Each tape element is represented as a block of n
bools. Whenever the machine
wants to read a symbol from the tape, it traverses over the block, performing n
branch
instructions to each any of the 2^n
results.
For the write
instruction, we have to use a goto
because we need to follow a different code
path depending on the local state, which is not available in the TM1 model, so instead we jump to
a label computed using the read value and the local state, which performs the writing and returns
to normal execution.
Emulation overhead is O(1)
. If not for the above write
behavior it would be 1-1 because we are
exploiting the 0-step behavior of regular commands to avoid taking steps, but there are
nevertheless a bounded number of write
calls between goto
statements because TM1 statements are
finitely long.
Equations
- Turing.TM1to1.instInhabitedΛ' = { default := Turing.TM1to1.Λ'.normal default }
Read a vector of length n
from the tape.
Equations
- One or more equations did not get rendered due to their size.
- Turing.TM1to1.readAux 0 f = f List.Vector.nil
Instances For
To read a symbol from the tape, we use readAux
to traverse the symbol,
then return to the original position with n
moves to the left.
Equations
- Turing.TM1to1.read dec f = Turing.TM1to1.readAux n fun (v : List.Vector Bool n) => Turing.TM1to1.move n Turing.Dir.left (f (dec v))
Instances For
Write a list of bools on the tape.
Equations
- Turing.TM1to1.write [] x✝ = x✝
- Turing.TM1to1.write (a :: l) x✝ = Turing.TM1.Stmt.write (fun (x : Bool) (x : σ) => a) (Turing.TM1.Stmt.move Turing.Dir.right (Turing.TM1to1.write l x✝))
Instances For
Translate a normal instruction. For the write
command, we use a goto
indirection so that
we can access the current value of the tape.
Equations
- One or more equations did not get rendered due to their size.
- Turing.TM1to1.trNormal dec (Turing.TM1.Stmt.move a q) = Turing.TM1to1.move n a (Turing.TM1to1.trNormal dec q)
- Turing.TM1to1.trNormal dec (Turing.TM1.Stmt.write a q) = Turing.TM1to1.read dec fun (a_1 : Γ) => Turing.TM1.Stmt.goto fun (x : Bool) (s : σ) => Turing.TM1to1.Λ'.write (a a_1 s) q
- Turing.TM1to1.trNormal dec (Turing.TM1.Stmt.load a q) = Turing.TM1to1.read dec fun (a_1 : Γ) => Turing.TM1.Stmt.load (fun (x : Bool) (s : σ) => a a_1 s) (Turing.TM1to1.trNormal dec q)
- Turing.TM1to1.trNormal dec (Turing.TM1.Stmt.goto l) = Turing.TM1to1.read dec fun (a : Γ) => Turing.TM1.Stmt.goto fun (x : Bool) (s : σ) => Turing.TM1to1.Λ'.normal (l a s)
- Turing.TM1to1.trNormal dec Turing.TM1.Stmt.halt = Turing.TM1.Stmt.halt
Instances For
The low level tape corresponding to the given tape over alphabet Γ
.
Equations
- Turing.TM1to1.trTape' enc0 L R = Turing.Tape.mk' (L.flatMap (fun (x : Γ) => (enc x).toList.reverse) ⋯) (R.flatMap (fun (x : Γ) => (enc x).toList) ⋯)
Instances For
The low level tape corresponding to the given tape over alphabet Γ
.
Equations
- Turing.TM1to1.trTape enc0 T = Turing.TM1to1.trTape' enc0 T.left T.right₀
Instances For
The top level program.
Equations
- Turing.TM1to1.tr enc dec M (Turing.TM1to1.Λ'.normal l) = Turing.TM1to1.trNormal dec (M l)
- Turing.TM1to1.tr enc dec M (Turing.TM1to1.Λ'.write a q) = Turing.TM1to1.write (enc a).toList (Turing.TM1to1.move n Turing.Dir.left (Turing.TM1to1.trNormal dec q))
Instances For
The machine configuration translation.
Equations
- Turing.TM1to1.trCfg enc enc0 { l := l, var := v, Tape := T } = { l := Option.map Turing.TM1to1.Λ'.normal l, var := v, Tape := Turing.TM1to1.trTape enc0 T }
Instances For
The set of accessible Λ'.write
machine states.
Equations
- Turing.TM1to1.writes (Turing.TM1.Stmt.move a q) = Turing.TM1to1.writes q
- Turing.TM1to1.writes (Turing.TM1.Stmt.write a q) = Finset.image (fun (a : Γ) => Turing.TM1to1.Λ'.write a q) Finset.univ ∪ Turing.TM1to1.writes q
- Turing.TM1to1.writes (Turing.TM1.Stmt.load a q) = Turing.TM1to1.writes q
- Turing.TM1to1.writes (Turing.TM1.Stmt.branch a q₁ q₂) = Turing.TM1to1.writes q₁ ∪ Turing.TM1to1.writes q₂
- Turing.TM1to1.writes (Turing.TM1.Stmt.goto l) = ∅
- Turing.TM1to1.writes Turing.TM1.Stmt.halt = ∅
Instances For
The set of accessible machine states, assuming that the input machine is supported on S
,
are the normal states embedded from S
, plus all write states accessible from these states.
Equations
- Turing.TM1to1.trSupp M S = S.biUnion fun (l : Λ) => insert (Turing.TM1to1.Λ'.normal l) (Turing.TM1to1.writes (M l))
Instances For
TM0 emulator in TM1 #
To establish that TM0 and TM1 are equivalent computational models, we must also have a TM0 emulator
in TM1. The main complication here is that TM0 allows an action to depend on the value at the head
and local state, while TM1 doesn't (in order to have more programming language-like semantics).
So we use a computed goto
to go to a state that performs the desired action and then returns to
normal execution.
One issue with this is that the halt
instruction is supposed to halt immediately, not take a step
to a halting state. To resolve this we do a check for halt
first, then goto
(with an
unreachable branch).
The machine states for a TM1 emulating a TM0 machine. States of the TM0 machine are embedded
as normal q
states, but the actual operation is split into two parts, a jump to act s q
followed by the action and a jump to the next normal
state.
- normal {Γ : Type u_1} {Λ : Type u_2} : Λ → Λ' Γ Λ
- act {Γ : Type u_1} {Λ : Type u_2} : TM0.Stmt Γ → Λ → Λ' Γ Λ
Instances For
Equations
- Turing.TM0to1.instInhabitedΛ' = { default := Turing.TM0to1.Λ'.normal default }
The program.
Equations
- One or more equations did not get rendered due to their size.
- Turing.TM0to1.tr M (Turing.TM0to1.Λ'.act (Turing.TM0.Stmt.move d) q) = Turing.TM1.Stmt.move d (Turing.TM1.Stmt.goto fun (x : Γ) (x : Unit) => Turing.TM0to1.Λ'.normal q)
Instances For
The configuration translation.
Equations
- Turing.TM0to1.trCfg M { q := q, Tape := T } = { l := bif (M q T.head).isSome then some (Turing.TM0to1.Λ'.normal q) else none, var := (), Tape := T }