Turing machine tapes #
This file defines the notion of a Turing machine tape, and the operations on it. A tape is a
bidirectional infinite sequence of cells, each of which stores an element of a given alphabet Γ
.
All but finitely many of the cells are required to hold the blank symbol default : Γ
.
Main definitions #
The BlankExtends
partial order holds of l₁
and l₂
if l₂
is obtained by adding
blanks (default : Γ
) to the end of l₁
.
Equations
- Turing.BlankExtends l₁ l₂ = ∃ (n : ℕ), l₂ = l₁ ++ List.replicate n default
Instances For
Any two extensions by blank l₁,l₂
of l
have a common join (which can be taken to be the
longer of l₁
and l₂
).
Instances For
BlankRel
is the symmetric closure of BlankExtends
, turning it into an equivalence
relation. Two lists are related by BlankRel
if one extends the other by blanks.
Equations
- Turing.BlankRel l₁ l₂ = (Turing.BlankExtends l₁ l₂ ∨ Turing.BlankExtends l₂ l₁)
Instances For
Given two BlankRel
lists, there exists (constructively) a common join.
Instances For
Given two BlankRel
lists, there exists (constructively) a common meet.
Instances For
Construct a setoid instance for BlankRel
.
Equations
- Turing.BlankRel.setoid Γ = { r := Turing.BlankRel, iseqv := ⋯ }
Instances For
Equations
- Turing.ListBlank.inhabited = { default := Quotient.mk'' [] }
Equations
- Turing.ListBlank.hasEmptyc = { emptyCollection := Quotient.mk'' [] }
A modified version of Quotient.liftOn'
specialized for ListBlank
, with the stronger
precondition BlankExtends
instead of BlankRel
.
Equations
- l.liftOn f H = Quotient.liftOn' l f ⋯
Instances For
The tail of a ListBlank
is well defined (up to the tail of blanks).
Equations
- l.tail = l.liftOn (fun (l : List Γ) => Turing.ListBlank.mk l.tail) ⋯
Instances For
We can cons an element onto a ListBlank
.
Equations
- Turing.ListBlank.cons a l = l.liftOn (fun (l : List Γ) => Turing.ListBlank.mk (a :: l)) ⋯
Instances For
Apply a function to a value stored at the nth position of the list.
Equations
- Turing.ListBlank.modifyNth f 0 x✝ = Turing.ListBlank.cons (f x✝.head) x✝.tail
- Turing.ListBlank.modifyNth f n.succ x✝ = Turing.ListBlank.cons x✝.head (Turing.ListBlank.modifyNth f n x✝.tail)
Instances For
Equations
- Turing.instInhabitedPointedMap = { default := { f := default, map_pt' := ⋯ } }
Equations
The map
function on lists is well defined on ListBlank
s provided that the map is
pointed.
Equations
- Turing.ListBlank.map f l = l.liftOn (fun (l : List Γ) => Turing.ListBlank.mk (List.map f.f l)) ⋯
Instances For
The i
-th projection as a pointed map.
Equations
- Turing.proj i = { f := fun (a : (i : ι) → Γ i) => a i, map_pt' := ⋯ }
Instances For
Append a list on the left side of a ListBlank
.
Equations
- Turing.ListBlank.append [] x✝ = x✝
- Turing.ListBlank.append (a :: l) x✝ = Turing.ListBlank.cons a (Turing.ListBlank.append l x✝)
Instances For
The flatMap
function on lists is well defined on ListBlank
s provided that the default
element is sent to a sequence of default elements.
Equations
- l.flatMap f hf = l.liftOn (fun (l : List Γ) => Turing.ListBlank.mk (l.flatMap f)) ⋯
Instances For
Alias of Turing.ListBlank.flatMap
.
The flatMap
function on lists is well defined on ListBlank
s provided that the default
element is sent to a sequence of default elements.
Instances For
Alias of Turing.ListBlank.flatMap_mk
.
Alias of Turing.ListBlank.cons_flatMap
.
The tape of a Turing machine is composed of a head element (which we imagine to be the
current position of the head), together with two ListBlank
s denoting the portions of the tape
going off to the left and right. When the Turing machine moves right, an element is pulled from the
right side and becomes the new head, while the head element is cons
ed onto the left side.
- head : Γ
The current position of the head.
- left : ListBlank Γ
The portion of the tape going off to the left.
- right : ListBlank Γ
The portion of the tape going off to the right.
Instances For
Equations
- Turing.instDecidableEqDir x✝ y✝ = if h : x✝.toCtorIdx = y✝.toCtorIdx then isTrue ⋯ else isFalse ⋯
Equations
- Turing.instInhabitedDir = { default := Turing.Dir.left }
The "inclusive" left side of the tape, including both left
and head
.
Equations
- T.left₀ = Turing.ListBlank.cons T.head T.left
Instances For
The "inclusive" right side of the tape, including both right
and head
.
Equations
- T.right₀ = Turing.ListBlank.cons T.head T.right
Instances For
Move the tape in response to a motion of the Turing machine. Note that T.move Dir.left
makes
T.left
smaller; the Turing machine is moving left and the tape is moving right.
Equations
- Turing.Tape.move Turing.Dir.left { head := a, left := L, right := R } = { head := L.head, left := L.tail, right := Turing.ListBlank.cons a R }
- Turing.Tape.move Turing.Dir.right { head := a, left := L, right := R } = { head := R.head, left := Turing.ListBlank.cons a L, right := R.tail }
Instances For
Construct a tape from a left side and an inclusive right side.
Equations
- Turing.Tape.mk' L R = { head := R.head, left := L, right := R.tail }
Instances For
Construct a tape from a left side and an inclusive right side.
Equations
Instances For
Construct a tape from a list, with the head of the list at the TM head and the rest going to the right.
Equations
- Turing.Tape.mk₁ l = Turing.Tape.mk₂ [] l
Instances For
Replace the current value of the head on the tape.
Equations
- Turing.Tape.write b T = { head := b, left := T.left, right := T.right }
Instances For
Apply a pointed map to a tape to change the alphabet.
Equations
- Turing.Tape.map f T = { head := f.f T.head, left := Turing.ListBlank.map f T.left, right := Turing.ListBlank.map f T.right }