Documentation

Mathlib.Computability.TuringMachine

Turing machines #

This file defines a sequence of simple machine languages, starting with Turing machines and working up to more complex languages based on Wang B-machines.

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:

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:

def Turing.BlankExtends {Γ : Type u_1} [Inhabited Γ] (l₁ : List Γ) (l₂ : List Γ) :

The BlankExtends partial order holds of l₁ and l₂ if l₂ is obtained by adding blanks (default : Γ) to the end of l₁.

Equations
Instances For
    theorem Turing.BlankExtends.trans {Γ : Type u_1} [Inhabited Γ] {l₁ : List Γ} {l₂ : List Γ} {l₃ : List Γ} :
    Turing.BlankExtends l₁ l₂Turing.BlankExtends l₂ l₃Turing.BlankExtends l₁ l₃
    theorem Turing.BlankExtends.below_of_le {Γ : Type u_1} [Inhabited Γ] {l : List Γ} {l₁ : List Γ} {l₂ : List Γ} :
    def Turing.BlankExtends.above {Γ : Type u_1} [Inhabited Γ] {l : List Γ} {l₁ : List Γ} {l₂ : List Γ} (h₁ : Turing.BlankExtends l l₁) (h₂ : Turing.BlankExtends l l₂) :
    { l' : List Γ // Turing.BlankExtends l₁ l' Turing.BlankExtends l₂ l' }

    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₂).

    Equations
    Instances For
      theorem Turing.BlankExtends.above_of_le {Γ : Type u_1} [Inhabited Γ] {l : List Γ} {l₁ : List Γ} {l₂ : List Γ} :
      def Turing.BlankRel {Γ : Type u_1} [Inhabited Γ] (l₁ : List Γ) (l₂ : List Γ) :

      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
      Instances For
        theorem Turing.BlankRel.refl {Γ : Type u_1} [Inhabited Γ] (l : List Γ) :
        theorem Turing.BlankRel.symm {Γ : Type u_1} [Inhabited Γ] {l₁ : List Γ} {l₂ : List Γ} :
        Turing.BlankRel l₁ l₂Turing.BlankRel l₂ l₁
        theorem Turing.BlankRel.trans {Γ : Type u_1} [Inhabited Γ] {l₁ : List Γ} {l₂ : List Γ} {l₃ : List Γ} :
        Turing.BlankRel l₁ l₂Turing.BlankRel l₂ l₃Turing.BlankRel l₁ l₃
        def Turing.BlankRel.above {Γ : Type u_1} [Inhabited Γ] {l₁ : List Γ} {l₂ : List Γ} (h : Turing.BlankRel l₁ l₂) :
        { l : List Γ // Turing.BlankExtends l₁ l Turing.BlankExtends l₂ l }

        Given two BlankRel lists, there exists (constructively) a common join.

        Equations
        Instances For
          def Turing.BlankRel.below {Γ : Type u_1} [Inhabited Γ] {l₁ : List Γ} {l₂ : List Γ} (h : Turing.BlankRel l₁ l₂) :
          { l : List Γ // Turing.BlankExtends l l₁ Turing.BlankExtends l l₂ }

          Given two BlankRel lists, there exists (constructively) a common meet.

          Equations
          Instances For
            theorem Turing.BlankRel.equivalence (Γ : Type u_1) [Inhabited Γ] :
            Equivalence Turing.BlankRel
            def Turing.BlankRel.setoid (Γ : Type u_1) [Inhabited Γ] :

            Construct a setoid instance for BlankRel.

            Equations
            Instances For
              def Turing.ListBlank (Γ : Type u_1) [Inhabited Γ] :
              Type u_1

              A ListBlank Γ is a quotient of List Γ by extension by blanks at the end. This is used to represent half-tapes of a Turing machine, so that we can pretend that the list continues infinitely with blanks.

              Equations
              Instances For
                Equations
                Equations
                @[reducible]
                def Turing.ListBlank.liftOn {Γ : Type u_1} [Inhabited Γ] {α : Sort u_2} (l : Turing.ListBlank Γ) (f : List Γα) (H : ∀ (a b : List Γ), Turing.BlankExtends a bf a = f b) :
                α

                A modified version of Quotient.liftOn' specialized for ListBlank, with the stronger precondition BlankExtends instead of BlankRel.

                Equations
                Instances For
                  def Turing.ListBlank.mk {Γ : Type u_1} [Inhabited Γ] :

                  The quotient map turning a List into a ListBlank.

                  Equations
                  • Turing.ListBlank.mk = Quotient.mk''
                  Instances For
                    theorem Turing.ListBlank.induction_on {Γ : Type u_1} [Inhabited Γ] {p : Turing.ListBlank ΓProp} (q : Turing.ListBlank Γ) (h : ∀ (a : List Γ), p (Turing.ListBlank.mk a)) :
                    p q
                    def Turing.ListBlank.head {Γ : Type u_1} [Inhabited Γ] (l : Turing.ListBlank Γ) :
                    Γ

                    The head of a ListBlank is well defined.

                    Equations
                    Instances For

                      The tail of a ListBlank is well defined (up to the tail of blanks).

                      Equations
                      Instances For
                        def Turing.ListBlank.cons {Γ : Type u_1} [Inhabited Γ] (a : Γ) (l : Turing.ListBlank Γ) :

                        We can cons an element onto a ListBlank.

                        Equations
                        Instances For
                          @[simp]

                          The cons and head/tail functions are mutually inverse, unlike in the case of List where this only holds for nonempty lists.

                          theorem Turing.ListBlank.exists_cons {Γ : Type u_1} [Inhabited Γ] (l : Turing.ListBlank Γ) :
                          ∃ (a : Γ) (l' : Turing.ListBlank Γ), l = Turing.ListBlank.cons a l'

                          The cons and head/tail functions are mutually inverse, unlike in the case of List where this only holds for nonempty lists.

                          def Turing.ListBlank.nth {Γ : Type u_1} [Inhabited Γ] (l : Turing.ListBlank Γ) (n : ) :
                          Γ

                          The n-th element of a ListBlank is well defined for all n : ℕ, unlike in a List.

                          Equations
                          Instances For
                            @[simp]
                            theorem Turing.ListBlank.ext {Γ : Type u_1} [i : Inhabited Γ] {L₁ : Turing.ListBlank Γ} {L₂ : Turing.ListBlank Γ} :
                            (∀ (i_1 : ), Turing.ListBlank.nth L₁ i_1 = Turing.ListBlank.nth L₂ i_1)L₁ = L₂
                            theorem Turing.ListBlank.nth_modifyNth {Γ : Type u_1} [Inhabited Γ] (f : ΓΓ) (n : ) (i : ) (L : Turing.ListBlank Γ) :
                            structure Turing.PointedMap (Γ : Type u) (Γ' : Type v) [Inhabited Γ] [Inhabited Γ'] :
                            Type (max u v)

                            A pointed map of Inhabited types is a map that sends one default value to the other.

                            • f : ΓΓ'
                            • map_pt' : self.f default = default
                            Instances For
                              instance Turing.instInhabitedPointedMap {Γ : Type u_1} {Γ' : Type u_2} [Inhabited Γ] [Inhabited Γ'] :
                              Equations
                              • Turing.instInhabitedPointedMap = { default := { f := default, map_pt' := } }
                              instance Turing.instCoeFunPointedMapForAll {Γ : Type u_1} {Γ' : Type u_2} [Inhabited Γ] [Inhabited Γ'] :
                              CoeFun (Turing.PointedMap Γ Γ') fun (x : Turing.PointedMap Γ Γ') => ΓΓ'
                              Equations
                              • Turing.instCoeFunPointedMapForAll = { coe := Turing.PointedMap.f }
                              theorem Turing.PointedMap.mk_val {Γ : Type u_1} {Γ' : Type u_2} [Inhabited Γ] [Inhabited Γ'] (f : ΓΓ') (pt : f default = default) :
                              { f := f, map_pt' := pt }.f = f
                              @[simp]
                              theorem Turing.PointedMap.map_pt {Γ : Type u_1} {Γ' : Type u_2} [Inhabited Γ] [Inhabited Γ'] (f : Turing.PointedMap Γ Γ') :
                              f.f default = default
                              @[simp]
                              theorem Turing.PointedMap.headI_map {Γ : Type u_1} {Γ' : Type u_2} [Inhabited Γ] [Inhabited Γ'] (f : Turing.PointedMap Γ Γ') (l : List Γ) :
                              def Turing.ListBlank.map {Γ : Type u_1} {Γ' : Type u_2} [Inhabited Γ] [Inhabited Γ'] (f : Turing.PointedMap Γ Γ') (l : Turing.ListBlank Γ) :

                              The map function on lists is well defined on ListBlanks provided that the map is pointed.

                              Equations
                              Instances For
                                @[simp]
                                @[simp]
                                theorem Turing.ListBlank.nth_map {Γ : Type u_1} {Γ' : Type u_2} [Inhabited Γ] [Inhabited Γ'] (f : Turing.PointedMap Γ Γ') (l : Turing.ListBlank Γ) (n : ) :
                                def Turing.proj {ι : Type u_1} {Γ : ιType u_2} [(i : ι) → Inhabited (Γ i)] (i : ι) :
                                Turing.PointedMap ((i : ι) → Γ i) (Γ i)

                                The i-th projection as a pointed map.

                                Equations
                                • Turing.proj i = { f := fun (a : (i : ι) → Γ i) => a i, map_pt' := }
                                Instances For
                                  theorem Turing.proj_map_nth {ι : Type u_1} {Γ : ιType u_2} [(i : ι) → Inhabited (Γ i)] (i : ι) (L : Turing.ListBlank ((i : ι) → Γ i)) (n : ) :
                                  theorem Turing.ListBlank.map_modifyNth {Γ : Type u_1} {Γ' : Type u_2} [Inhabited Γ] [Inhabited Γ'] (F : Turing.PointedMap Γ Γ') (f : ΓΓ) (f' : Γ'Γ') (H : ∀ (x : Γ), F.f (f x) = f' (F.f x)) (n : ) (L : Turing.ListBlank Γ) :

                                  Append a list on the left side of a ListBlank.

                                  Equations
                                  Instances For
                                    @[simp]
                                    theorem Turing.ListBlank.append_mk {Γ : Type u_1} [Inhabited Γ] (l₁ : List Γ) (l₂ : List Γ) :
                                    theorem Turing.ListBlank.append_assoc {Γ : Type u_1} [Inhabited Γ] (l₁ : List Γ) (l₂ : List Γ) (l₃ : Turing.ListBlank Γ) :
                                    def Turing.ListBlank.bind {Γ : Type u_1} {Γ' : Type u_2} [Inhabited Γ] [Inhabited Γ'] (l : Turing.ListBlank Γ) (f : ΓList Γ') (hf : ∃ (n : ), f default = List.replicate n default) :

                                    The bind function on lists is well defined on ListBlanks provided that the default element is sent to a sequence of default elements.

                                    Equations
                                    Instances For
                                      @[simp]
                                      theorem Turing.ListBlank.bind_mk {Γ : Type u_1} {Γ' : Type u_2} [Inhabited Γ] [Inhabited Γ'] (l : List Γ) (f : ΓList Γ') (hf : ∃ (n : ), f default = List.replicate n default) :
                                      @[simp]
                                      theorem Turing.ListBlank.cons_bind {Γ : Type u_1} {Γ' : Type u_2} [Inhabited Γ] [Inhabited Γ'] (a : Γ) (l : Turing.ListBlank Γ) (f : ΓList Γ') (hf : ∃ (n : ), f default = List.replicate n default) :
                                      structure Turing.Tape (Γ : Type u_1) [Inhabited Γ] :
                                      Type u_1

                                      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 ListBlanks 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 consed onto the left side.

                                      Instances For
                                        Equations
                                        • Turing.Tape.inhabited = { default := { head := default, left := default, right := default } }
                                        inductive Turing.Dir :

                                        A direction for the Turing machine move command, either left or right.

                                        Instances For

                                          The "inclusive" left side of the tape, including both left and head.

                                          Equations
                                          Instances For

                                            The "inclusive" right side of the tape, including both right and head.

                                            Equations
                                            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
                                              • One or more equations did not get rendered due to their size.
                                              Instances For

                                                Construct a tape from a left side and an inclusive right side.

                                                Equations
                                                Instances For
                                                  @[simp]
                                                  theorem Turing.Tape.mk'_left {Γ : Type u_1} [Inhabited Γ] (L : Turing.ListBlank Γ) (R : Turing.ListBlank Γ) :
                                                  (Turing.Tape.mk' L R).left = L
                                                  theorem Turing.Tape.exists_mk' {Γ : Type u_1} [Inhabited Γ] (T : Turing.Tape Γ) :
                                                  ∃ (L : Turing.ListBlank Γ) (R : Turing.ListBlank Γ), T = Turing.Tape.mk' L R
                                                  def Turing.Tape.mk₂ {Γ : Type u_1} [Inhabited Γ] (L : List Γ) (R : List Γ) :

                                                  Construct a tape from a left side and an inclusive right side.

                                                  Equations
                                                  Instances For
                                                    def Turing.Tape.mk₁ {Γ : Type u_1} [Inhabited Γ] (l : List Γ) :

                                                    Construct a tape from a list, with the head of the list at the TM head and the rest going to the right.

                                                    Equations
                                                    Instances For
                                                      def Turing.Tape.nth {Γ : Type u_1} [Inhabited Γ] (T : Turing.Tape Γ) :
                                                      Γ

                                                      The nth function of a tape is integer-valued, with index 0 being the head, negative indexes on the left and positive indexes on the right. (Picture a number line.)

                                                      Equations
                                                      Instances For
                                                        @[simp]
                                                        theorem Turing.Tape.nth_zero {Γ : Type u_1} [Inhabited Γ] (T : Turing.Tape Γ) :
                                                        Turing.Tape.nth T 0 = T.head
                                                        def Turing.Tape.write {Γ : Type u_1} [Inhabited Γ] (b : Γ) (T : Turing.Tape Γ) :

                                                        Replace the current value of the head on the tape.

                                                        Equations
                                                        Instances For
                                                          @[simp]
                                                          theorem Turing.Tape.write_self {Γ : Type u_1} [Inhabited Γ] (T : Turing.Tape Γ) :
                                                          Turing.Tape.write T.head T = T
                                                          @[simp]
                                                          theorem Turing.Tape.write_nth {Γ : Type u_1} [Inhabited Γ] (b : Γ) (T : Turing.Tape Γ) {i : } :
                                                          Turing.Tape.nth (Turing.Tape.write b T) i = if i = 0 then b else Turing.Tape.nth T i
                                                          def Turing.Tape.map {Γ : Type u_1} {Γ' : Type u_2} [Inhabited Γ] [Inhabited Γ'] (f : Turing.PointedMap Γ Γ') (T : Turing.Tape Γ) :

                                                          Apply a pointed map to a tape to change the alphabet.

                                                          Equations
                                                          Instances For
                                                            @[simp]
                                                            theorem Turing.Tape.map_fst {Γ : Type u_1} {Γ' : Type u_2} [Inhabited Γ] [Inhabited Γ'] (f : Turing.PointedMap Γ Γ') (T : Turing.Tape Γ) :
                                                            (Turing.Tape.map f T).head = f.f T.head
                                                            @[simp]
                                                            theorem Turing.Tape.map_write {Γ : Type u_1} {Γ' : Type u_2} [Inhabited Γ] [Inhabited Γ'] (f : Turing.PointedMap Γ Γ') (b : Γ) (T : Turing.Tape Γ) :
                                                            theorem Turing.Tape.map_mk₂ {Γ : Type u_1} {Γ' : Type u_2} [Inhabited Γ] [Inhabited Γ'] (f : Turing.PointedMap Γ Γ') (L : List Γ) (R : List Γ) :
                                                            theorem Turing.Tape.map_mk₁ {Γ : Type u_1} {Γ' : Type u_2} [Inhabited Γ] [Inhabited Γ'] (f : Turing.PointedMap Γ Γ') (l : List Γ) :
                                                            def Turing.eval {σ : Type u_1} (f : σOption σ) :
                                                            σPart σ

                                                            Run a state transition function σ → Option σ "to completion". The return value is the last state returned before a none result. If the state transition function always returns some, then the computation diverges, returning Part.none.

                                                            Equations
                                                            Instances For
                                                              def Turing.Reaches {σ : Type u_1} (f : σOption σ) :
                                                              σσProp

                                                              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
                                                              Instances For
                                                                def Turing.Reaches₁ {σ : Type u_1} (f : σOption σ) :
                                                                σσProp

                                                                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
                                                                Instances For
                                                                  theorem Turing.reaches₁_eq {σ : Type u_1} {f : σOption σ} {a : σ} {b : σ} {c : σ} (h : f a = f b) :
                                                                  theorem Turing.reaches_total {σ : Type u_1} {f : σOption σ} {a : σ} {b : σ} {c : σ} (hab : Turing.Reaches f a b) (hac : Turing.Reaches f a c) :
                                                                  theorem Turing.reaches₁_fwd {σ : Type u_1} {f : σOption σ} {a : σ} {b : σ} {c : σ} (h₁ : Turing.Reaches₁ f a c) (h₂ : b f a) :
                                                                  def Turing.Reaches₀ {σ : Type u_1} (f : σOption σ) (a : σ) (b : σ) :

                                                                  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
                                                                  Instances For
                                                                    theorem Turing.Reaches₀.trans {σ : Type u_1} {f : σOption σ} {a : σ} {b : σ} {c : σ} (h₁ : Turing.Reaches₀ f a b) (h₂ : Turing.Reaches₀ f b c) :
                                                                    theorem Turing.Reaches₀.refl {σ : Type u_1} {f : σOption σ} (a : σ) :
                                                                    theorem Turing.Reaches₀.single {σ : Type u_1} {f : σOption σ} {a : σ} {b : σ} (h : b f a) :
                                                                    theorem Turing.Reaches₀.head {σ : Type u_1} {f : σOption σ} {a : σ} {b : σ} {c : σ} (h : b f a) (h₂ : Turing.Reaches₀ f b c) :
                                                                    theorem Turing.Reaches₀.tail {σ : Type u_1} {f : σOption σ} {a : σ} {b : σ} {c : σ} (h₁ : Turing.Reaches₀ f a b) (h : c f b) :
                                                                    theorem Turing.reaches₀_eq {σ : Type u_1} {f : σOption σ} {a : σ} {b : σ} (e : f a = f b) :
                                                                    theorem Turing.Reaches₁.to₀ {σ : Type u_1} {f : σOption σ} {a : σ} {b : σ} (h : Turing.Reaches₁ f a b) :
                                                                    theorem Turing.Reaches.to₀ {σ : Type u_1} {f : σOption σ} {a : σ} {b : σ} (h : Turing.Reaches f a b) :
                                                                    theorem Turing.Reaches₀.tail' {σ : Type u_1} {f : σOption σ} {a : σ} {b : σ} {c : σ} (h : Turing.Reaches₀ f a b) (h₂ : c f b) :
                                                                    def Turing.evalInduction {σ : Type u_2} {f : σOption σ} {b : σ} {C : σSort u_1} {a : σ} (h : b Turing.eval f a) (H : (a : σ) → b Turing.eval f a((a' : σ) → f a = some a'C a')C a) :
                                                                    C a

                                                                    (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
                                                                      theorem Turing.mem_eval {σ : Type u_1} {f : σOption σ} {a : σ} {b : σ} :
                                                                      b Turing.eval f a Turing.Reaches f a b f b = none
                                                                      theorem Turing.eval_maximal₁ {σ : Type u_1} {f : σOption σ} {a : σ} {b : σ} (h : b Turing.eval f a) (c : σ) :
                                                                      theorem Turing.eval_maximal {σ : Type u_1} {f : σOption σ} {a : σ} {b : σ} (h : b Turing.eval f a) {c : σ} :
                                                                      Turing.Reaches f b c c = b
                                                                      theorem Turing.reaches_eval {σ : Type u_1} {f : σOption σ} {a : σ} {b : σ} (ab : Turing.Reaches f a b) :
                                                                      def Turing.Respects {σ₁ : Type u_1} {σ₂ : Type u_2} (f₁ : σ₁Option σ₁) (f₂ : σ₂Option σ₂) (tr : σ₁σ₂Prop) :

                                                                      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
                                                                        theorem Turing.tr_reaches₁ {σ₁ : Type u_1} {σ₂ : Type u_2} {f₁ : σ₁Option σ₁} {f₂ : σ₂Option σ₂} {tr : σ₁σ₂Prop} (H : Turing.Respects f₁ f₂ tr) {a₁ : σ₁} {a₂ : σ₂} (aa : tr a₁ a₂) {b₁ : σ₁} (ab : Turing.Reaches₁ f₁ a₁ b₁) :
                                                                        ∃ (b₂ : σ₂), tr b₁ b₂ Turing.Reaches₁ f₂ a₂ b₂
                                                                        theorem Turing.tr_reaches {σ₁ : Type u_1} {σ₂ : Type u_2} {f₁ : σ₁Option σ₁} {f₂ : σ₂Option σ₂} {tr : σ₁σ₂Prop} (H : Turing.Respects f₁ f₂ tr) {a₁ : σ₁} {a₂ : σ₂} (aa : tr a₁ a₂) {b₁ : σ₁} (ab : Turing.Reaches f₁ a₁ b₁) :
                                                                        ∃ (b₂ : σ₂), tr b₁ b₂ Turing.Reaches f₂ a₂ b₂
                                                                        theorem Turing.tr_reaches_rev {σ₁ : Type u_1} {σ₂ : Type u_2} {f₁ : σ₁Option σ₁} {f₂ : σ₂Option σ₂} {tr : σ₁σ₂Prop} (H : Turing.Respects f₁ f₂ tr) {a₁ : σ₁} {a₂ : σ₂} (aa : tr a₁ a₂) {b₂ : σ₂} (ab : Turing.Reaches f₂ a₂ b₂) :
                                                                        ∃ (c₁ : σ₁) (c₂ : σ₂), Turing.Reaches f₂ b₂ c₂ tr c₁ c₂ Turing.Reaches f₁ a₁ c₁
                                                                        theorem Turing.tr_eval {σ₁ : Type u_1} {σ₂ : Type u_2} {f₁ : σ₁Option σ₁} {f₂ : σ₂Option σ₂} {tr : σ₁σ₂Prop} (H : Turing.Respects f₁ f₂ tr) {a₁ : σ₁} {b₁ : σ₁} {a₂ : σ₂} (aa : tr a₁ a₂) (ab : b₁ Turing.eval f₁ a₁) :
                                                                        ∃ (b₂ : σ₂), tr b₁ b₂ b₂ Turing.eval f₂ a₂
                                                                        theorem Turing.tr_eval_rev {σ₁ : Type u_1} {σ₂ : Type u_2} {f₁ : σ₁Option σ₁} {f₂ : σ₂Option σ₂} {tr : σ₁σ₂Prop} (H : Turing.Respects f₁ f₂ tr) {a₁ : σ₁} {b₂ : σ₂} {a₂ : σ₂} (aa : tr a₁ a₂) (ab : b₂ Turing.eval f₂ a₂) :
                                                                        ∃ (b₁ : σ₁), tr b₁ b₂ b₁ Turing.eval f₁ a₁
                                                                        theorem Turing.tr_eval_dom {σ₁ : Type u_1} {σ₂ : Type u_2} {f₁ : σ₁Option σ₁} {f₂ : σ₂Option σ₂} {tr : σ₁σ₂Prop} (H : Turing.Respects f₁ f₂ tr) {a₁ : σ₁} {a₂ : σ₂} (aa : tr a₁ a₂) :
                                                                        (Turing.eval f₂ a₂).Dom (Turing.eval f₁ a₁).Dom
                                                                        def Turing.FRespects {σ₁ : Type u_1} {σ₂ : Type u_2} (f₂ : σ₂Option σ₂) (tr : σ₁σ₂) (a₂ : σ₂) :
                                                                        Option σ₁Prop

                                                                        A simpler version of Respects when the state transition relation tr is a function.

                                                                        Equations
                                                                        Instances For
                                                                          theorem Turing.frespects_eq {σ₁ : Type u_1} {σ₂ : Type u_2} {f₂ : σ₂Option σ₂} {tr : σ₁σ₂} {a₂ : σ₂} {b₂ : σ₂} (h : f₂ a₂ = f₂ b₂) {b₁ : Option σ₁} :
                                                                          Turing.FRespects f₂ tr a₂ b₁ Turing.FRespects f₂ tr b₂ b₁
                                                                          theorem Turing.fun_respects {σ₁ : Type u_1} {σ₂ : Type u_2} {f₁ : σ₁Option σ₁} {f₂ : σ₂Option σ₂} {tr : σ₁σ₂} :
                                                                          (Turing.Respects f₁ f₂ fun (a : σ₁) (b : σ₂) => tr a = b) ∀ ⦃a₁ : σ₁⦄, Turing.FRespects f₂ tr (tr a₁) (f₁ a₁)
                                                                          theorem Turing.tr_eval' {σ₁ : Type u_1} {σ₂ : Type u_1} (f₁ : σ₁Option σ₁) (f₂ : σ₂Option σ₂) (tr : σ₁σ₂) (H : Turing.Respects f₁ f₂ fun (a : σ₁) (b : σ₂) => tr a = b) (a₁ : σ₁) :
                                                                          Turing.eval f₂ (tr a₁) = tr <$> Turing.eval f₁ a₁

                                                                          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:

                                                                          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.)

                                                                          inductive Turing.TM0.Stmt (Γ : Type u_1) :
                                                                          Type u_1

                                                                          A Turing machine "statement" is just a command to either move left or right, or write a symbol on the tape.

                                                                          Instances For
                                                                            def Turing.TM0.Machine (Γ : Type u_1) (Λ : Type u_2) [Inhabited Λ] :
                                                                            Type (max (max u_1 u_2) u_1 u_2)

                                                                            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
                                                                            Instances For
                                                                              instance Turing.TM0.Machine.inhabited (Γ : Type u_1) (Λ : Type u_2) [Inhabited Λ] :
                                                                              Equations
                                                                              structure Turing.TM0.Cfg (Γ : Type u_1) [Inhabited Γ] (Λ : Type u_2) :
                                                                              Type (max u_1 u_2)

                                                                              The configuration state of a Turing machine during operation consists of a label (machine state), and a tape, 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.

                                                                              Instances For
                                                                                instance Turing.TM0.Cfg.inhabited (Γ : Type u_1) [Inhabited Γ] (Λ : Type u_2) [Inhabited Λ] :
                                                                                Equations
                                                                                def Turing.TM0.step {Γ : Type u_1} [Inhabited Γ] {Λ : Type u_2} [Inhabited Λ] (M : Turing.TM0.Machine Γ Λ) :

                                                                                Execution semantics of the Turing machine.

                                                                                Equations
                                                                                • One or more equations did not get rendered due to their size.
                                                                                Instances For
                                                                                  def Turing.TM0.Reaches {Γ : Type u_1} [Inhabited Γ] {Λ : Type u_2} [Inhabited Λ] (M : Turing.TM0.Machine Γ Λ) :
                                                                                  Turing.TM0.Cfg Γ ΛTuring.TM0.Cfg Γ ΛProp

                                                                                  The statement Reaches M s₁ s₂ means that s₂ is obtained starting from s₁ after a finite number of steps from s₂.

                                                                                  Equations
                                                                                  Instances For
                                                                                    def Turing.TM0.init {Γ : Type u_1} [Inhabited Γ] {Λ : Type u_2} [Inhabited Λ] (l : List Γ) :

                                                                                    The initial configuration.

                                                                                    Equations
                                                                                    Instances For
                                                                                      def Turing.TM0.eval {Γ : Type u_1} [Inhabited Γ] {Λ : Type u_2} [Inhabited Λ] (M : Turing.TM0.Machine Γ Λ) (l : List Γ) :

                                                                                      Evaluate a Turing machine on initial input to a final state, if it terminates.

                                                                                      Equations
                                                                                      Instances For
                                                                                        def Turing.TM0.Supports {Γ : Type u_1} {Λ : Type u_2} [Inhabited Λ] (M : Turing.TM0.Machine Γ Λ) (S : Set Λ) :

                                                                                        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
                                                                                          theorem Turing.TM0.step_supports {Γ : Type u_1} [Inhabited Γ] {Λ : Type u_2} [Inhabited Λ] (M : Turing.TM0.Machine Γ Λ) {S : Set Λ} (ss : Turing.TM0.Supports M S) {c : Turing.TM0.Cfg Γ Λ} {c' : Turing.TM0.Cfg Γ Λ} :
                                                                                          c' Turing.TM0.step M cc.q Sc'.q S
                                                                                          theorem Turing.TM0.univ_supports {Γ : Type u_1} {Λ : Type u_2} [Inhabited Λ] (M : Turing.TM0.Machine Γ Λ) :
                                                                                          def Turing.TM0.Stmt.map {Γ : Type u_1} [Inhabited Γ] {Γ' : Type u_2} [Inhabited Γ'] (f : Turing.PointedMap Γ Γ') :

                                                                                          Map a TM statement across a function. This does nothing to move statements and maps the write values.

                                                                                          Equations
                                                                                          Instances For
                                                                                            def Turing.TM0.Cfg.map {Γ : Type u_1} [Inhabited Γ] {Γ' : Type u_2} [Inhabited Γ'] {Λ : Type u_3} {Λ' : Type u_4} (f : Turing.PointedMap Γ Γ') (g : ΛΛ') :
                                                                                            Turing.TM0.Cfg Γ ΛTuring.TM0.Cfg Γ' Λ'

                                                                                            Map a configuration across a function, given f : Γ → Γ' a map of the alphabets and g : Λ → Λ' a map of the machine states.

                                                                                            Equations
                                                                                            Instances For
                                                                                              def Turing.TM0.Machine.map {Γ : Type u_1} [Inhabited Γ] {Γ' : Type u_2} [Inhabited Γ'] {Λ : Type u_3} [Inhabited Λ] {Λ' : Type u_4} [Inhabited Λ'] (M : Turing.TM0.Machine Γ Λ) (f₁ : Turing.PointedMap Γ Γ') (f₂ : Turing.PointedMap Γ' Γ) (g₁ : ΛΛ') (g₂ : Λ'Λ) :

                                                                                              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
                                                                                              Instances For
                                                                                                theorem Turing.TM0.Machine.map_step {Γ : Type u_1} [Inhabited Γ] {Γ' : Type u_2} [Inhabited Γ'] {Λ : Type u_3} [Inhabited Λ] {Λ' : Type u_4} [Inhabited Λ'] (M : Turing.TM0.Machine Γ Λ) (f₁ : Turing.PointedMap Γ Γ') (f₂ : Turing.PointedMap Γ' Γ) (g₁ : ΛΛ') (g₂ : Λ'Λ) {S : Set Λ} (f₂₁ : Function.RightInverse f₁.f f₂.f) (g₂₁ : qS, g₂ (g₁ q) = q) (c : Turing.TM0.Cfg Γ Λ) :
                                                                                                c.q SOption.map (Turing.TM0.Cfg.map f₁ g₁) (Turing.TM0.step M c) = Turing.TM0.step (Turing.TM0.Machine.map M f₁ f₂ g₁ g₂) (Turing.TM0.Cfg.map f₁ g₁ c)
                                                                                                theorem Turing.TM0.map_init {Γ : Type u_1} [Inhabited Γ] {Γ' : Type u_2} [Inhabited Γ'] {Λ : Type u_3} [Inhabited Λ] {Λ' : Type u_4} [Inhabited Λ'] (f₁ : Turing.PointedMap Γ Γ') (g₁ : Turing.PointedMap Λ Λ') (l : List Γ) :
                                                                                                theorem Turing.TM0.Machine.map_respects {Γ : Type u_1} [Inhabited Γ] {Γ' : Type u_2} [Inhabited Γ'] {Λ : Type u_3} [Inhabited Λ] {Λ' : Type u_4} [Inhabited Λ'] (M : Turing.TM0.Machine Γ Λ) (f₁ : Turing.PointedMap Γ Γ') (f₂ : Turing.PointedMap Γ' Γ) (g₁ : Turing.PointedMap Λ Λ') (g₂ : Λ'Λ) {S : Set Λ} (ss : Turing.TM0.Supports M S) (f₂₁ : Function.RightInverse f₁.f f₂.f) (g₂₁ : qS, g₂ (g₁.f q) = q) :
                                                                                                Turing.Respects (Turing.TM0.step M) (Turing.TM0.step (Turing.TM0.Machine.map M f₁ f₂ g₁.f g₂)) fun (a : Turing.TM0.Cfg Γ Λ) (b : Turing.TM0.Cfg Γ' Λ') => a.q S Turing.TM0.Cfg.map f₁ g₁.f a = b

                                                                                                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 Stmts can be as follows:

                                                                                                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.

                                                                                                inductive Turing.TM1.Stmt (Γ : Type u_1) (Λ : Type u_2) (σ : Type u_3) :
                                                                                                Type (max (max u_1 u_2) u_3)

                                                                                                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.

                                                                                                Instances For
                                                                                                  instance Turing.TM1.Stmt.inhabited (Γ : Type u_1) (Λ : Type u_2) (σ : Type u_3) :
                                                                                                  Equations
                                                                                                  structure Turing.TM1.Cfg (Γ : Type u_1) [Inhabited Γ] (Λ : Type u_2) (σ : Type u_3) :
                                                                                                  Type (max (max u_1 u_2) u_3)

                                                                                                  The configuration of a TM1 machine is given by the currently evaluating statement, the variable store value, and the tape.

                                                                                                  Instances For
                                                                                                    instance Turing.TM1.Cfg.inhabited (Γ : Type u_1) [Inhabited Γ] (Λ : Type u_2) (σ : Type u_3) [Inhabited σ] :
                                                                                                    Equations
                                                                                                    def Turing.TM1.stepAux {Γ : Type u_1} [Inhabited Γ] {Λ : Type u_2} {σ : Type u_3} :
                                                                                                    Turing.TM1.Stmt Γ Λ σσTuring.Tape ΓTuring.TM1.Cfg Γ Λ σ

                                                                                                    The semantics of TM1 evaluation.

                                                                                                    Equations
                                                                                                    Instances For
                                                                                                      def Turing.TM1.step {Γ : Type u_1} [Inhabited Γ] {Λ : Type u_2} {σ : Type u_3} (M : ΛTuring.TM1.Stmt Γ Λ σ) :
                                                                                                      Turing.TM1.Cfg Γ Λ σOption (Turing.TM1.Cfg Γ Λ σ)

                                                                                                      The state transition function.

                                                                                                      Equations
                                                                                                      Instances For
                                                                                                        def Turing.TM1.SupportsStmt {Γ : Type u_1} {Λ : Type u_2} {σ : Type u_3} (S : Finset Λ) :
                                                                                                        Turing.TM1.Stmt Γ Λ σProp

                                                                                                        A set S of labels supports the statement q if all the goto statements in q refer only to other functions in S.

                                                                                                        Equations
                                                                                                        Instances For
                                                                                                          theorem Turing.TM1.stmts₁_self {Γ : Type u_1} {Λ : Type u_2} {σ : Type u_3} {q : Turing.TM1.Stmt Γ Λ σ} :
                                                                                                          theorem Turing.TM1.stmts₁_trans {Γ : Type u_1} {Λ : Type u_2} {σ : Type u_3} {q₁ : Turing.TM1.Stmt Γ Λ σ} {q₂ : Turing.TM1.Stmt Γ Λ σ} :
                                                                                                          theorem Turing.TM1.stmts₁_supportsStmt_mono {Γ : Type u_1} {Λ : Type u_2} {σ : Type u_3} {S : Finset Λ} {q₁ : Turing.TM1.Stmt Γ Λ σ} {q₂ : Turing.TM1.Stmt Γ Λ σ} (h : q₁ Turing.TM1.stmts₁ q₂) (hs : Turing.TM1.SupportsStmt S q₂) :
                                                                                                          noncomputable def Turing.TM1.stmts {Γ : Type u_1} {Λ : Type u_2} {σ : Type u_3} (M : ΛTuring.TM1.Stmt Γ Λ σ) (S : Finset Λ) :

                                                                                                          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
                                                                                                          Instances For
                                                                                                            theorem Turing.TM1.stmts_trans {Γ : Type u_1} {Λ : Type u_2} {σ : Type u_3} {M : ΛTuring.TM1.Stmt Γ Λ σ} {S : Finset Λ} {q₁ : Turing.TM1.Stmt Γ Λ σ} {q₂ : Turing.TM1.Stmt Γ Λ σ} (h₁ : q₁ Turing.TM1.stmts₁ q₂) :
                                                                                                            def Turing.TM1.Supports {Γ : Type u_1} {Λ : Type u_2} {σ : Type u_3} [Inhabited Λ] (M : ΛTuring.TM1.Stmt Γ Λ σ) (S : Finset Λ) :

                                                                                                            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
                                                                                                            Instances For
                                                                                                              theorem Turing.TM1.stmts_supportsStmt {Γ : Type u_1} {Λ : Type u_2} {σ : Type u_3} [Inhabited Λ] {M : ΛTuring.TM1.Stmt Γ Λ σ} {S : Finset Λ} {q : Turing.TM1.Stmt Γ Λ σ} (ss : Turing.TM1.Supports M S) :
                                                                                                              theorem Turing.TM1.step_supports {Γ : Type u_1} [Inhabited Γ] {Λ : Type u_2} {σ : Type u_3} [Inhabited Λ] (M : ΛTuring.TM1.Stmt Γ Λ σ) {S : Finset Λ} (ss : Turing.TM1.Supports M S) {c : Turing.TM1.Cfg Γ Λ σ} {c' : Turing.TM1.Cfg Γ Λ σ} :
                                                                                                              c' Turing.TM1.step M cc.l Finset.insertNone Sc'.l Finset.insertNone S
                                                                                                              def Turing.TM1.init {Γ : Type u_1} [Inhabited Γ] {Λ : Type u_2} {σ : Type u_3} [Inhabited Λ] [Inhabited σ] (l : List Γ) :

                                                                                                              The initial state, given a finite input that is placed on the tape starting at the TM head and going to the right.

                                                                                                              Equations
                                                                                                              Instances For
                                                                                                                def Turing.TM1.eval {Γ : Type u_1} [Inhabited Γ] {Λ : Type u_2} {σ : Type u_3} [Inhabited Λ] [Inhabited σ] (M : ΛTuring.TM1.Stmt Γ Λ σ) (l : List Γ) :

                                                                                                                Evaluate a TM to completion, resulting in an output list on the tape (with an indeterminate number of blanks on the end).

                                                                                                                Equations
                                                                                                                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:

                                                                                                                  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.

                                                                                                                  def Turing.TM1to0.Λ' {Γ : Type u_1} {Λ : Type u_2} {σ : Type u_3} (M : ΛTuring.TM1.Stmt Γ Λ σ) :
                                                                                                                  Type (max (max (max u_3 u_2) u_1) u_3)

                                                                                                                  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
                                                                                                                  Instances For
                                                                                                                    instance Turing.TM1to0.instInhabitedΛ' {Γ : Type u_1} {Λ : Type u_2} [Inhabited Λ] {σ : Type u_3} [Inhabited σ] (M : ΛTuring.TM1.Stmt Γ Λ σ) :
                                                                                                                    Equations
                                                                                                                    def Turing.TM1to0.trAux {Γ : Type u_1} {Λ : Type u_2} {σ : Type u_3} (M : ΛTuring.TM1.Stmt Γ Λ σ) (s : Γ) :

                                                                                                                    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
                                                                                                                    Instances For
                                                                                                                      def Turing.TM1to0.tr {Γ : Type u_1} {Λ : Type u_2} [Inhabited Λ] {σ : Type u_3} [Inhabited σ] (M : ΛTuring.TM1.Stmt Γ Λ σ) :

                                                                                                                      The translated TM0 machine (given the TM1 machine input).

                                                                                                                      Equations
                                                                                                                      Instances For
                                                                                                                        def Turing.TM1to0.trCfg {Γ : Type u_1} [Inhabited Γ] {Λ : Type u_2} {σ : Type u_3} (M : ΛTuring.TM1.Stmt Γ Λ σ) :

                                                                                                                        Translate configurations from TM1 to TM0.

                                                                                                                        Equations
                                                                                                                        Instances For
                                                                                                                          theorem Turing.TM1to0.tr_respects {Γ : Type u_1} [Inhabited Γ] {Λ : Type u_2} [Inhabited Λ] {σ : Type u_3} [Inhabited σ] (M : ΛTuring.TM1.Stmt Γ Λ σ) :
                                                                                                                          theorem Turing.TM1to0.tr_eval {Γ : Type u_1} [Inhabited Γ] {Λ : Type u_2} [Inhabited Λ] {σ : Type u_3} [Inhabited σ] (M : ΛTuring.TM1.Stmt Γ Λ σ) (l : List Γ) :
                                                                                                                          noncomputable def Turing.TM1to0.trStmts {Γ : Type u_1} {Λ : Type u_2} {σ : Type u_3} (M : ΛTuring.TM1.Stmt Γ Λ σ) [Fintype σ] (S : Finset Λ) :

                                                                                                                          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
                                                                                                                            theorem Turing.TM1to0.tr_supports {Γ : Type u_1} {Λ : Type u_2} [Inhabited Λ] {σ : Type u_3} [Inhabited σ] (M : ΛTuring.TM1.Stmt Γ Λ σ) [Fintype σ] {S : Finset Λ} (ss : Turing.TM1.Supports M S) :

                                                                                                                            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.

                                                                                                                            theorem Turing.TM1to1.exists_enc_dec {Γ : Type u_1} [Inhabited Γ] [Finite Γ] :
                                                                                                                            ∃ (n : ) (enc : ΓVector Bool n) (dec : Vector Bool nΓ), enc default = Vector.replicate n false ∀ (a : Γ), dec (enc a) = a
                                                                                                                            inductive Turing.TM1to1.Λ' {Γ : Type u_1} {Λ : Type u_2} {σ : Type u_3} :
                                                                                                                            Type (max (max u_1 u_2) u_3)

                                                                                                                            The configuration state of the TM.

                                                                                                                            • normal: {Γ : Type u_1} → {Λ : Type u_2} → {σ : Type u_3} → ΛTuring.TM1to1.Λ'
                                                                                                                            • write: {Γ : Type u_1} → {Λ : Type u_2} → {σ : Type u_3} → ΓTuring.TM1.Stmt Γ Λ σTuring.TM1to1.Λ'
                                                                                                                            Instances For
                                                                                                                              instance Turing.TM1to1.instInhabitedΛ' {Γ : Type u_1} {Λ : Type u_2} [Inhabited Λ] {σ : Type u_3} :
                                                                                                                              Inhabited Turing.TM1to1.Λ'
                                                                                                                              Equations
                                                                                                                              def Turing.TM1to1.readAux {Γ : Type u_1} {Λ : Type u_2} {σ : Type u_3} (n : ) :
                                                                                                                              (Vector Bool nTuring.TM1.Stmt Bool Turing.TM1to1.Λ' σ)Turing.TM1.Stmt Bool Turing.TM1to1.Λ' σ

                                                                                                                              Read a vector of length n from the tape.

                                                                                                                              Equations
                                                                                                                              Instances For
                                                                                                                                def Turing.TM1to1.move {Γ : Type u_1} {Λ : Type u_2} {σ : Type u_3} {n : } (d : Turing.Dir) (q : Turing.TM1.Stmt Bool Turing.TM1to1.Λ' σ) :
                                                                                                                                Turing.TM1.Stmt Bool Turing.TM1to1.Λ' σ

                                                                                                                                A move left or right corresponds to n moves across the super-cell.

                                                                                                                                Equations
                                                                                                                                Instances For
                                                                                                                                  def Turing.TM1to1.read {Γ : Type u_1} {Λ : Type u_2} {σ : Type u_3} {n : } (dec : Vector Bool nΓ) (f : ΓTuring.TM1.Stmt Bool Turing.TM1to1.Λ' σ) :
                                                                                                                                  Turing.TM1.Stmt Bool Turing.TM1to1.Λ' σ

                                                                                                                                  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
                                                                                                                                  Instances For
                                                                                                                                    def Turing.TM1to1.write {Γ : Type u_1} {Λ : Type u_2} {σ : Type u_3} :
                                                                                                                                    List BoolTuring.TM1.Stmt Bool Turing.TM1to1.Λ' σTuring.TM1.Stmt Bool Turing.TM1to1.Λ' σ

                                                                                                                                    Write a list of bools on the tape.

                                                                                                                                    Equations
                                                                                                                                    Instances For
                                                                                                                                      def Turing.TM1to1.trNormal {Γ : Type u_1} {Λ : Type u_2} {σ : Type u_3} {n : } (dec : Vector Bool nΓ) :
                                                                                                                                      Turing.TM1.Stmt Γ Λ σTuring.TM1.Stmt Bool Turing.TM1to1.Λ' σ

                                                                                                                                      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
                                                                                                                                      Instances For
                                                                                                                                        theorem Turing.TM1to1.stepAux_move {Γ : Type u_1} {Λ : Type u_2} {σ : Type u_3} {n : } (d : Turing.Dir) (q : Turing.TM1.Stmt Bool Turing.TM1to1.Λ' σ) (v : σ) (T : Turing.Tape Bool) :
                                                                                                                                        theorem Turing.TM1to1.supportsStmt_move {Γ : Type u_1} {Λ : Type u_2} {σ : Type u_3} {n : } {S : Finset Turing.TM1to1.Λ'} {d : Turing.Dir} {q : Turing.TM1.Stmt Bool Turing.TM1to1.Λ' σ} :
                                                                                                                                        theorem Turing.TM1to1.supportsStmt_write {Γ : Type u_1} {Λ : Type u_2} {σ : Type u_3} {S : Finset Turing.TM1to1.Λ'} {l : List Bool} {q : Turing.TM1.Stmt Bool Turing.TM1to1.Λ' σ} :
                                                                                                                                        theorem Turing.TM1to1.supportsStmt_read {Γ : Type u_1} {Λ : Type u_2} {σ : Type u_3} {n : } (dec : Vector Bool nΓ) {S : Finset Turing.TM1to1.Λ'} {f : ΓTuring.TM1.Stmt Bool Turing.TM1to1.Λ' σ} :
                                                                                                                                        def Turing.TM1to1.trTape' {Γ : Type u_1} [Inhabited Γ] {n : } {enc : ΓVector Bool n} (enc0 : enc default = Vector.replicate n false) (L : Turing.ListBlank Γ) (R : Turing.ListBlank Γ) :

                                                                                                                                        The low level tape corresponding to the given tape over alphabet Γ.

                                                                                                                                        Equations
                                                                                                                                        • One or more equations did not get rendered due to their size.
                                                                                                                                        Instances For
                                                                                                                                          def Turing.TM1to1.trTape {Γ : Type u_1} [Inhabited Γ] {n : } {enc : ΓVector Bool n} (enc0 : enc default = Vector.replicate n false) (T : Turing.Tape Γ) :

                                                                                                                                          The low level tape corresponding to the given tape over alphabet Γ.

                                                                                                                                          Equations
                                                                                                                                          Instances For
                                                                                                                                            theorem Turing.TM1to1.trTape_mk' {Γ : Type u_1} [Inhabited Γ] {n : } {enc : ΓVector Bool n} (enc0 : enc default = Vector.replicate n false) (L : Turing.ListBlank Γ) (R : Turing.ListBlank Γ) :
                                                                                                                                            def Turing.TM1to1.tr {Γ : Type u_1} {Λ : Type u_2} {σ : Type u_3} {n : } (enc : ΓVector Bool n) (dec : Vector Bool nΓ) (M : ΛTuring.TM1.Stmt Γ Λ σ) :
                                                                                                                                            Turing.TM1to1.Λ'Turing.TM1.Stmt Bool Turing.TM1to1.Λ' σ

                                                                                                                                            The top level program.

                                                                                                                                            Equations
                                                                                                                                            • One or more equations did not get rendered due to their size.
                                                                                                                                            Instances For
                                                                                                                                              def Turing.TM1to1.trCfg {Γ : Type u_1} [Inhabited Γ] {Λ : Type u_2} {σ : Type u_3} {n : } (enc : ΓVector Bool n) (enc0 : enc default = Vector.replicate n false) :
                                                                                                                                              Turing.TM1.Cfg Γ Λ σTuring.TM1.Cfg Bool Turing.TM1to1.Λ' σ

                                                                                                                                              The machine configuration translation.

                                                                                                                                              Equations
                                                                                                                                              Instances For
                                                                                                                                                theorem Turing.TM1to1.stepAux_write {Γ : Type u_1} [Inhabited Γ] {Λ : Type u_2} {σ : Type u_3} {n : } {enc : ΓVector Bool n} (enc0 : enc default = Vector.replicate n false) (q : Turing.TM1.Stmt Bool Turing.TM1to1.Λ' σ) (v : σ) (a : Γ) (b : Γ) (L : Turing.ListBlank Γ) (R : Turing.ListBlank Γ) :
                                                                                                                                                theorem Turing.TM1to1.stepAux_read {Γ : Type u_1} [Inhabited Γ] {Λ : Type u_2} {σ : Type u_3} {n : } {enc : ΓVector Bool n} (dec : Vector Bool nΓ) (enc0 : enc default = Vector.replicate n false) (encdec : ∀ (a : Γ), dec (enc a) = a) (f : ΓTuring.TM1.Stmt Bool Turing.TM1to1.Λ' σ) (v : σ) (L : Turing.ListBlank Γ) (R : Turing.ListBlank Γ) :
                                                                                                                                                theorem Turing.TM1to1.tr_respects {Γ : Type u_1} [Inhabited Γ] {Λ : Type u_2} {σ : Type u_3} {n : } {enc : ΓVector Bool n} (dec : Vector Bool nΓ) (enc0 : enc default = Vector.replicate n false) (M : ΛTuring.TM1.Stmt Γ Λ σ) (encdec : ∀ (a : Γ), dec (enc a) = a) {enc₀ : enc default = Vector.replicate n false} :
                                                                                                                                                Turing.Respects (Turing.TM1.step M) (Turing.TM1.step (Turing.TM1to1.tr enc dec M)) fun (c₁ : Turing.TM1.Cfg Γ Λ σ) (c₂ : Turing.TM1.Cfg Bool Turing.TM1to1.Λ' σ) => Turing.TM1to1.trCfg enc enc₀ c₁ = c₂
                                                                                                                                                noncomputable def Turing.TM1to1.trSupp {Γ : Type u_1} {Λ : Type u_2} {σ : Type u_3} (M : ΛTuring.TM1.Stmt Γ Λ σ) [Fintype Γ] (S : Finset Λ) :
                                                                                                                                                Finset Turing.TM1to1.Λ'

                                                                                                                                                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
                                                                                                                                                Instances For
                                                                                                                                                  theorem Turing.TM1to1.tr_supports {Γ : Type u_1} {Λ : Type u_2} [Inhabited Λ] {σ : Type u_3} {n : } {enc : ΓVector Bool n} (dec : Vector Bool nΓ) (M : ΛTuring.TM1.Stmt Γ Λ σ) [Fintype Γ] {S : Finset Λ} (ss : Turing.TM1.Supports M S) :

                                                                                                                                                  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).

                                                                                                                                                  inductive Turing.TM0to1.Λ' {Γ : Type u_1} {Λ : Type u_2} :
                                                                                                                                                  Type (max u_1 u_2)

                                                                                                                                                  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} → ΛTuring.TM0to1.Λ'
                                                                                                                                                  • act: {Γ : Type u_1} → {Λ : Type u_2} → Turing.TM0.Stmt ΓΛTuring.TM0to1.Λ'
                                                                                                                                                  Instances For
                                                                                                                                                    instance Turing.TM0to1.instInhabitedΛ' {Γ : Type u_1} {Λ : Type u_2} [Inhabited Λ] :
                                                                                                                                                    Inhabited Turing.TM0to1.Λ'
                                                                                                                                                    Equations
                                                                                                                                                    def Turing.TM0to1.tr {Γ : Type u_1} {Λ : Type u_2} [Inhabited Λ] (M : Turing.TM0.Machine Γ Λ) :
                                                                                                                                                    Turing.TM0to1.Λ'Turing.TM1.Stmt Γ Turing.TM0to1.Λ' Unit

                                                                                                                                                    The program.

                                                                                                                                                    Equations
                                                                                                                                                    • One or more equations did not get rendered due to their size.
                                                                                                                                                    Instances For
                                                                                                                                                      def Turing.TM0to1.trCfg {Γ : Type u_1} [Inhabited Γ] {Λ : Type u_2} [Inhabited Λ] (M : Turing.TM0.Machine Γ Λ) :
                                                                                                                                                      Turing.TM0.Cfg Γ ΛTuring.TM1.Cfg Γ Turing.TM0to1.Λ' Unit

                                                                                                                                                      The configuration translation.

                                                                                                                                                      Equations
                                                                                                                                                      Instances For
                                                                                                                                                        theorem Turing.TM0to1.tr_respects {Γ : Type u_1} [Inhabited Γ] {Λ : Type u_2} [Inhabited Λ] (M : Turing.TM0.Machine Γ Λ) :

                                                                                                                                                        The TM2 model #

                                                                                                                                                        The TM2 model removes the tape entirely from the TM1 model, replacing it with an arbitrary (finite) collection of stacks, each with elements of different types (the alphabet of stack k : K is Γ k). The statements are:

                                                                                                                                                        The configuration is a tuple (l, var, stk) where l : Option Λ is the current label to run or none for the halting state, var : σ is the (finite) internal state, and stk : ∀ k, List (Γ k) is the collection of stacks. (Note that unlike the TM0 and TM1 models, these are not ListBlanks, they have definite ends that can be detected by the pop command.)

                                                                                                                                                        Given a designated stack k and a value L : List (Γ k), the initial configuration has all the stacks empty except the designated "input" stack; in eval this designated stack also functions as the output stack.

                                                                                                                                                        inductive Turing.TM2.Stmt {K : Type u_1} (Γ : KType u_2) (Λ : Type u_3) (σ : Type u_4) :
                                                                                                                                                        Type (max (max (max u_1 u_2) u_3) u_4)

                                                                                                                                                        The TM2 model removes the tape entirely from the TM1 model, replacing it with an arbitrary (finite) collection of stacks. The operation push puts an element on one of the stacks, and pop removes an element from a stack (and modifying the internal state based on the result). peek modifies the internal state but does not remove an element.

                                                                                                                                                        Instances For
                                                                                                                                                          instance Turing.TM2.Stmt.inhabited {K : Type u_1} (Γ : KType u_2) (Λ : Type u_3) (σ : Type u_4) :
                                                                                                                                                          Equations
                                                                                                                                                          structure Turing.TM2.Cfg {K : Type u_1} (Γ : KType u_2) (Λ : Type u_3) (σ : Type u_4) :
                                                                                                                                                          Type (max (max (max u_1 u_2) u_3) u_4)

                                                                                                                                                          A configuration in the TM2 model is a label (or none for the halt state), the state of local variables, and the stacks. (Note that the stacks are not ListBlanks, they have a definite size.)

                                                                                                                                                          • l : Option Λ
                                                                                                                                                          • var : σ
                                                                                                                                                          • stk : (k : K) → List (Γ k)
                                                                                                                                                          Instances For
                                                                                                                                                            instance Turing.TM2.Cfg.inhabited {K : Type u_1} (Γ : KType u_2) (Λ : Type u_3) (σ : Type u_4) [Inhabited σ] :
                                                                                                                                                            Equations
                                                                                                                                                            def Turing.TM2.stepAux {K : Type u_1} [DecidableEq K] {Γ : KType u_2} {Λ : Type u_3} {σ : Type u_4} :
                                                                                                                                                            Turing.TM2.Stmt Γ Λ σσ((k : K) → List (Γ k))Turing.TM2.Cfg Γ Λ σ

                                                                                                                                                            The step function for the TM2 model.

                                                                                                                                                            Equations
                                                                                                                                                            Instances For
                                                                                                                                                              def Turing.TM2.step {K : Type u_1} [DecidableEq K] {Γ : KType u_2} {Λ : Type u_3} {σ : Type u_4} (M : ΛTuring.TM2.Stmt Γ Λ σ) :
                                                                                                                                                              Turing.TM2.Cfg Γ Λ σOption (Turing.TM2.Cfg Γ Λ σ)

                                                                                                                                                              The step function for the TM2 model.

                                                                                                                                                              Equations
                                                                                                                                                              Instances For
                                                                                                                                                                def Turing.TM2.Reaches {K : Type u_1} [DecidableEq K] {Γ : KType u_2} {Λ : Type u_3} {σ : Type u_4} (M : ΛTuring.TM2.Stmt Γ Λ σ) :
                                                                                                                                                                Turing.TM2.Cfg Γ Λ σTuring.TM2.Cfg Γ Λ σProp

                                                                                                                                                                The (reflexive) reachability relation for the TM2 model.

                                                                                                                                                                Equations
                                                                                                                                                                Instances For
                                                                                                                                                                  theorem Turing.TM2.stmts₁_self {K : Type u_1} {Γ : KType u_2} {Λ : Type u_3} {σ : Type u_4} {q : Turing.TM2.Stmt Γ Λ σ} :
                                                                                                                                                                  theorem Turing.TM2.stmts₁_trans {K : Type u_1} {Γ : KType u_2} {Λ : Type u_3} {σ : Type u_4} {q₁ : Turing.TM2.Stmt Γ Λ σ} {q₂ : Turing.TM2.Stmt Γ Λ σ} :
                                                                                                                                                                  theorem Turing.TM2.stmts₁_supportsStmt_mono {K : Type u_1} {Γ : KType u_2} {Λ : Type u_3} {σ : Type u_4} {S : Finset Λ} {q₁ : Turing.TM2.Stmt Γ Λ σ} {q₂ : Turing.TM2.Stmt Γ Λ σ} (h : q₁ Turing.TM2.stmts₁ q₂) (hs : Turing.TM2.SupportsStmt S q₂) :
                                                                                                                                                                  noncomputable def Turing.TM2.stmts {K : Type u_1} {Γ : KType u_2} {Λ : Type u_3} {σ : Type u_4} (M : ΛTuring.TM2.Stmt Γ Λ σ) (S : Finset Λ) :

                                                                                                                                                                  The set of statements accessible from initial set S of labels.

                                                                                                                                                                  Equations
                                                                                                                                                                  Instances For
                                                                                                                                                                    theorem Turing.TM2.stmts_trans {K : Type u_1} {Γ : KType u_2} {Λ : Type u_3} {σ : Type u_4} {M : ΛTuring.TM2.Stmt Γ Λ σ} {S : Finset Λ} {q₁ : Turing.TM2.Stmt Γ Λ σ} {q₂ : Turing.TM2.Stmt Γ Λ σ} (h₁ : q₁ Turing.TM2.stmts₁ q₂) :
                                                                                                                                                                    def Turing.TM2.Supports {K : Type u_1} {Γ : KType u_2} {Λ : Type u_3} {σ : Type u_4} [Inhabited Λ] (M : ΛTuring.TM2.Stmt Γ Λ σ) (S : Finset Λ) :

                                                                                                                                                                    Given a TM2 machine M and a set S of states, Supports M S means that all states in S jump only to other states in S.

                                                                                                                                                                    Equations
                                                                                                                                                                    Instances For
                                                                                                                                                                      theorem Turing.TM2.stmts_supportsStmt {K : Type u_1} {Γ : KType u_2} {Λ : Type u_3} {σ : Type u_4} [Inhabited Λ] {M : ΛTuring.TM2.Stmt Γ Λ σ} {S : Finset Λ} {q : Turing.TM2.Stmt Γ Λ σ} (ss : Turing.TM2.Supports M S) :
                                                                                                                                                                      theorem Turing.TM2.step_supports {K : Type u_1} [DecidableEq K] {Γ : KType u_2} {Λ : Type u_3} {σ : Type u_4} [Inhabited Λ] (M : ΛTuring.TM2.Stmt Γ Λ σ) {S : Finset Λ} (ss : Turing.TM2.Supports M S) {c : Turing.TM2.Cfg Γ Λ σ} {c' : Turing.TM2.Cfg Γ Λ σ} :
                                                                                                                                                                      c' Turing.TM2.step M cc.l Finset.insertNone Sc'.l Finset.insertNone S
                                                                                                                                                                      def Turing.TM2.init {K : Type u_1} [DecidableEq K] {Γ : KType u_2} {Λ : Type u_3} {σ : Type u_4} [Inhabited Λ] [Inhabited σ] (k : K) (L : List (Γ k)) :

                                                                                                                                                                      The initial state of the TM2 model. The input is provided on a designated stack.

                                                                                                                                                                      Equations
                                                                                                                                                                      Instances For
                                                                                                                                                                        def Turing.TM2.eval {K : Type u_1} [DecidableEq K] {Γ : KType u_2} {Λ : Type u_3} {σ : Type u_4} [Inhabited Λ] [Inhabited σ] (M : ΛTuring.TM2.Stmt Γ Λ σ) (k : K) (L : List (Γ k)) :
                                                                                                                                                                        Part (List (Γ k))

                                                                                                                                                                        Evaluates a TM2 program to completion, with the output on the same stack as the input.

                                                                                                                                                                        Equations
                                                                                                                                                                        Instances For

                                                                                                                                                                          TM2 emulator in TM1 #

                                                                                                                                                                          To prove that TM2 computable functions are TM1 computable, we need to reduce each TM2 program to a TM1 program. So suppose a TM2 program is given. This program has to maintain a whole collection of stacks, but we have only one tape, so we must "multiplex" them all together. Pictorially, if stack 1 contains [a, b] and stack 2 contains [c, d, e, f] then the tape looks like this:

                                                                                                                                                                           bottom:  ... | _ | T | _ | _ | _ | _ | ...
                                                                                                                                                                           stack 1: ... | _ | b | a | _ | _ | _ | ...
                                                                                                                                                                           stack 2: ... | _ | f | e | d | c | _ | ...
                                                                                                                                                                          

                                                                                                                                                                          where a tape element is a vertical slice through the diagram. Here the alphabet is Γ' := Bool × ∀ k, Option (Γ k), where:

                                                                                                                                                                          In "resting" position, the TM is sitting at the position marked bottom. For non-stack actions, it operates in place, but for the stack actions push, peek, and pop, it must shuttle to the end of the appropriate stack, make its changes, and then return to the bottom. So the states are:

                                                                                                                                                                          Because of the shuttling, emulation overhead is O(n), where n is the current maximum of the length of all stacks. Therefore a program that takes k steps to run in TM2 takes O((m+k)k) steps to run when emulated in TM1, where m is the length of the input.

                                                                                                                                                                          theorem Turing.TM2to1.stk_nth_val {K : Type u_1} {Γ : KType u_2} {L : Turing.ListBlank ((k : K) → Option (Γ k))} {k : K} {S : List (Γ k)} (n : ) (hL : Turing.ListBlank.map (Turing.proj k) L = Turing.ListBlank.mk (List.reverse (List.map some S))) :
                                                                                                                                                                          def Turing.TM2to1.Γ' {K : Type u_1} {Γ : KType u_2} :
                                                                                                                                                                          Type (max u_1 u_2)

                                                                                                                                                                          The alphabet of the TM2 simulator on TM1 is a marker for the stack bottom, plus a vector of stack elements for each stack, or none if the stack does not extend this far.

                                                                                                                                                                          Equations
                                                                                                                                                                          Instances For
                                                                                                                                                                            instance Turing.TM2to1.Γ'.inhabited {K : Type u_1} {Γ : KType u_2} :
                                                                                                                                                                            Inhabited Turing.TM2to1.Γ'
                                                                                                                                                                            Equations
                                                                                                                                                                            • Turing.TM2to1.Γ'.inhabited = { default := (false, fun (x : K) => none) }
                                                                                                                                                                            instance Turing.TM2to1.Γ'.fintype {K : Type u_1} [DecidableEq K] {Γ : KType u_2} [Fintype K] [(k : K) → Fintype (Γ k)] :
                                                                                                                                                                            Fintype Turing.TM2to1.Γ'
                                                                                                                                                                            Equations
                                                                                                                                                                            def Turing.TM2to1.addBottom {K : Type u_1} {Γ : KType u_2} (L : Turing.ListBlank ((k : K) → Option (Γ k))) :
                                                                                                                                                                            Turing.ListBlank Turing.TM2to1.Γ'

                                                                                                                                                                            The bottom marker is fixed throughout the calculation, so we use the addBottom function to express the program state in terms of a tape with only the stacks themselves.

                                                                                                                                                                            Equations
                                                                                                                                                                            Instances For
                                                                                                                                                                              theorem Turing.TM2to1.addBottom_map {K : Type u_1} {Γ : KType u_2} (L : Turing.ListBlank ((k : K) → Option (Γ k))) :
                                                                                                                                                                              Turing.ListBlank.map { f := Prod.snd, map_pt' := } (Turing.TM2to1.addBottom L) = L
                                                                                                                                                                              theorem Turing.TM2to1.addBottom_modifyNth {K : Type u_1} {Γ : KType u_2} (f : ((k : K) → Option (Γ k))(k : K) → Option (Γ k)) (L : Turing.ListBlank ((k : K) → Option (Γ k))) (n : ) :
                                                                                                                                                                              theorem Turing.TM2to1.addBottom_nth_snd {K : Type u_1} {Γ : KType u_2} (L : Turing.ListBlank ((k : K) → Option (Γ k))) (n : ) :
                                                                                                                                                                              theorem Turing.TM2to1.addBottom_nth_succ_fst {K : Type u_1} {Γ : KType u_2} (L : Turing.ListBlank ((k : K) → Option (Γ k))) (n : ) :
                                                                                                                                                                              theorem Turing.TM2to1.addBottom_head_fst {K : Type u_1} {Γ : KType u_2} (L : Turing.ListBlank ((k : K) → Option (Γ k))) :
                                                                                                                                                                              inductive Turing.TM2to1.StAct {K : Type u_1} {Γ : KType u_2} {σ : Type u_4} (k : K) :
                                                                                                                                                                              Type (max u_2 u_4)

                                                                                                                                                                              A stack action is a command that interacts with the top of a stack. Our default position is at the bottom of all the stacks, so we have to hold on to this action while going to the end to modify the stack.

                                                                                                                                                                              Instances For
                                                                                                                                                                                instance Turing.TM2to1.StAct.inhabited {K : Type u_1} {Γ : KType u_2} {σ : Type u_4} {k : K} :
                                                                                                                                                                                Equations
                                                                                                                                                                                def Turing.TM2to1.stRun {K : Type u_1} {Γ : KType u_2} {Λ : Type u_3} {σ : Type u_4} {k : K} :

                                                                                                                                                                                The TM2 statement corresponding to a stack action.

                                                                                                                                                                                Equations
                                                                                                                                                                                • One or more equations did not get rendered due to their size.
                                                                                                                                                                                Instances For
                                                                                                                                                                                  def Turing.TM2to1.stVar {K : Type u_1} {Γ : KType u_2} {σ : Type u_4} {k : K} (v : σ) (l : List (Γ k)) :

                                                                                                                                                                                  The effect of a stack action on the local variables, given the value of the stack.

                                                                                                                                                                                  Equations
                                                                                                                                                                                  Instances For
                                                                                                                                                                                    def Turing.TM2to1.stWrite {K : Type u_1} {Γ : KType u_2} {σ : Type u_4} {k : K} (v : σ) (l : List (Γ k)) :

                                                                                                                                                                                    The effect of a stack action on the stack.

                                                                                                                                                                                    Equations
                                                                                                                                                                                    Instances For
                                                                                                                                                                                      def Turing.TM2to1.stmtStRec {K : Type u_1} {Γ : KType u_2} {Λ : Type u_3} {σ : Type u_4} {C : Turing.TM2.Stmt Γ Λ σSort l} (H₁ : (k : K) → (s : Turing.TM2to1.StAct k) → (q : Turing.TM2.Stmt Γ Λ σ) → C qC (Turing.TM2to1.stRun s q)) (H₂ : (a : σσ) → (q : Turing.TM2.Stmt Γ Λ σ) → C qC (Turing.TM2.Stmt.load a q)) (H₃ : (p : σBool) → (q₁ q₂ : Turing.TM2.Stmt Γ Λ σ) → C q₁C q₂C (Turing.TM2.Stmt.branch p q₁ q₂)) (H₄ : (l : σΛ) → C (Turing.TM2.Stmt.goto l)) (H₅ : C Turing.TM2.Stmt.halt) (n : Turing.TM2.Stmt Γ Λ σ) :
                                                                                                                                                                                      C n

                                                                                                                                                                                      We have partitioned the TM2 statements into "stack actions", which require going to the end of the stack, and all other actions, which do not. This is a modified recursor which lumps the stack actions into one.

                                                                                                                                                                                      Equations
                                                                                                                                                                                      Instances For
                                                                                                                                                                                        theorem Turing.TM2to1.supports_run {K : Type u_1} {Γ : KType u_2} {Λ : Type u_3} {σ : Type u_4} (S : Finset Λ) {k : K} (s : Turing.TM2to1.StAct k) (q : Turing.TM2.Stmt Γ Λ σ) :
                                                                                                                                                                                        inductive Turing.TM2to1.Λ' {K : Type u_1} {Γ : KType u_2} {Λ : Type u_3} {σ : Type u_4} :
                                                                                                                                                                                        Type (max (max (max u_1 u_2) u_3) u_4)

                                                                                                                                                                                        The machine states of the TM2 emulator. We can either be in a normal state when waiting for the next TM2 action, or we can be in the "go" and "return" states to go to the top of the stack and return to the bottom, respectively.

                                                                                                                                                                                        Instances For
                                                                                                                                                                                          instance Turing.TM2to1.Λ'.inhabited {K : Type u_1} {Γ : KType u_2} {Λ : Type u_3} [Inhabited Λ] {σ : Type u_4} :
                                                                                                                                                                                          Inhabited Turing.TM2to1.Λ'
                                                                                                                                                                                          Equations
                                                                                                                                                                                          def Turing.TM2to1.trStAct {K : Type u_1} [DecidableEq K] {Γ : KType u_2} {Λ : Type u_3} {σ : Type u_4} {k : K} (q : Turing.TM1.Stmt Turing.TM2to1.Γ' Turing.TM2to1.Λ' σ) :
                                                                                                                                                                                          Turing.TM2to1.StAct kTuring.TM1.Stmt Turing.TM2to1.Γ' Turing.TM2to1.Λ' σ

                                                                                                                                                                                          The program corresponding to state transitions at the end of a stack. Here we start out just after the top of the stack, and should end just after the new top of the stack.

                                                                                                                                                                                          Equations
                                                                                                                                                                                          • One or more equations did not get rendered due to their size.
                                                                                                                                                                                          Instances For
                                                                                                                                                                                            def Turing.TM2to1.trInit {K : Type u_1} [DecidableEq K] {Γ : KType u_2} (k : K) (L : List (Γ k)) :
                                                                                                                                                                                            List Turing.TM2to1.Γ'

                                                                                                                                                                                            The initial state for the TM2 emulator, given an initial TM2 state. All stacks start out empty except for the input stack, and the stack bottom mark is set at the head.

                                                                                                                                                                                            Equations
                                                                                                                                                                                            Instances For
                                                                                                                                                                                              theorem Turing.TM2to1.step_run {K : Type u_1} [DecidableEq K] {Γ : KType u_2} {Λ : Type u_3} {σ : Type u_4} {k : K} (q : Turing.TM2.Stmt Γ Λ σ) (v : σ) (S : (k : K) → List (Γ k)) (s : Turing.TM2to1.StAct k) :
                                                                                                                                                                                              def Turing.TM2to1.trNormal {K : Type u_1} {Γ : KType u_2} {Λ : Type u_3} {σ : Type u_4} :
                                                                                                                                                                                              Turing.TM2.Stmt Γ Λ σTuring.TM1.Stmt Turing.TM2to1.Γ' Turing.TM2to1.Λ' σ

                                                                                                                                                                                              The translation of TM2 statements to TM1 statements. regular actions have direct equivalents, but stack actions are deferred by going to the corresponding go state, so that we can find the appropriate stack top.

                                                                                                                                                                                              Equations
                                                                                                                                                                                              Instances For
                                                                                                                                                                                                theorem Turing.TM2to1.trNormal_run {K : Type u_1} {Γ : KType u_2} {Λ : Type u_3} {σ : Type u_4} {k : K} (s : Turing.TM2to1.StAct k) (q : Turing.TM2.Stmt Γ Λ σ) :
                                                                                                                                                                                                Turing.TM2to1.trNormal (Turing.TM2to1.stRun s q) = Turing.TM1.Stmt.goto fun (x : Turing.TM2to1.Γ') (x : σ) => Turing.TM2to1.Λ'.go k s q
                                                                                                                                                                                                theorem Turing.TM2to1.tr_respects_aux₂ {K : Type u_1} [DecidableEq K] {Γ : KType u_2} {Λ : Type u_3} {σ : Type u_4} {k : K} {q : Turing.TM1.Stmt Turing.TM2to1.Γ' Turing.TM2to1.Λ' σ} {v : σ} {S : (k : K) → List (Γ k)} {L : Turing.ListBlank ((k : K) → Option (Γ k))} (hL : ∀ (k : K), Turing.ListBlank.map (Turing.proj k) L = Turing.ListBlank.mk (List.reverse (List.map some (S k)))) (o : Turing.TM2to1.StAct k) :
                                                                                                                                                                                                def Turing.TM2to1.tr {K : Type u_1} [DecidableEq K] {Γ : KType u_2} {Λ : Type u_3} {σ : Type u_4} (M : ΛTuring.TM2.Stmt Γ Λ σ) :
                                                                                                                                                                                                Turing.TM2to1.Λ'Turing.TM1.Stmt Turing.TM2to1.Γ' Turing.TM2to1.Λ' σ

                                                                                                                                                                                                The TM2 emulator machine states written as a TM1 program. This handles the go and ret states, which shuttle to and from a stack top.

                                                                                                                                                                                                Equations
                                                                                                                                                                                                • One or more equations did not get rendered due to their size.
                                                                                                                                                                                                Instances For
                                                                                                                                                                                                  inductive Turing.TM2to1.TrCfg {K : Type u_1} {Γ : KType u_2} {Λ : Type u_3} {σ : Type u_4} :
                                                                                                                                                                                                  Turing.TM2.Cfg Γ Λ σTuring.TM1.Cfg Turing.TM2to1.Γ' Turing.TM2to1.Λ' σProp

                                                                                                                                                                                                  The relation between TM2 configurations and TM1 configurations of the TM2 emulator.

                                                                                                                                                                                                  Instances For
                                                                                                                                                                                                    theorem Turing.TM2to1.tr_respects_aux₁ {K : Type u_1} [DecidableEq K] {Γ : KType u_2} {Λ : Type u_3} {σ : Type u_4} (M : ΛTuring.TM2.Stmt Γ Λ σ) {k : K} (o : Turing.TM2to1.StAct k) (q : Turing.TM2.Stmt Γ Λ σ) (v : σ) {S : List (Γ k)} {L : Turing.ListBlank ((k : K) → Option (Γ k))} (hL : Turing.ListBlank.map (Turing.proj k) L = Turing.ListBlank.mk (List.reverse (List.map some S))) (n : ) (H : n List.length S) :
                                                                                                                                                                                                    theorem Turing.TM2to1.tr_respects_aux₃ {K : Type u_1} [DecidableEq K] {Γ : KType u_2} {Λ : Type u_3} {σ : Type u_4} (M : ΛTuring.TM2.Stmt Γ Λ σ) {q : Turing.TM2.Stmt Γ Λ σ} {v : σ} {L : Turing.ListBlank ((k : K) → Option (Γ k))} (n : ) :
                                                                                                                                                                                                    theorem Turing.TM2to1.tr_respects_aux {K : Type u_1} [DecidableEq K] {Γ : KType u_2} {Λ : Type u_3} {σ : Type u_4} (M : ΛTuring.TM2.Stmt Γ Λ σ) {q : Turing.TM2.Stmt (fun (k : K) => Γ k) Λ σ} {v : σ} {T : Turing.ListBlank ((i : K) → Option (Γ i))} {k : K} {S : (k : K) → List (Γ k)} (hT : ∀ (k : K), Turing.ListBlank.map (Turing.proj k) T = Turing.ListBlank.mk (List.reverse (List.map some (S k)))) (o : Turing.TM2to1.StAct k) (IH : ∀ {v : σ} {S : (k : K) → List (Γ k)} {T : Turing.ListBlank ((k : K) → Option (Γ k))}, (∀ (k : K), Turing.ListBlank.map (Turing.proj k) T = Turing.ListBlank.mk (List.reverse (List.map some (S k))))∃ (b : Turing.TM1.Cfg Turing.TM2to1.Γ' Turing.TM2to1.Λ' σ), Turing.TM2to1.TrCfg (Turing.TM2.stepAux q v S) b Turing.Reaches (Turing.TM1.step (Turing.TM2to1.tr M)) (Turing.TM1.stepAux (Turing.TM2to1.trNormal q) v (Turing.Tape.mk' (Turing.TM2to1.addBottom T))) b) :
                                                                                                                                                                                                    theorem Turing.TM2to1.tr_respects {K : Type u_1} [DecidableEq K] {Γ : KType u_2} {Λ : Type u_3} {σ : Type u_4} (M : ΛTuring.TM2.Stmt Γ Λ σ) :
                                                                                                                                                                                                    theorem Turing.TM2to1.trCfg_init {K : Type u_1} [DecidableEq K] {Γ : KType u_2} {Λ : Type u_3} [Inhabited Λ] {σ : Type u_4} [Inhabited σ] (k : K) (L : List (Γ k)) :
                                                                                                                                                                                                    theorem Turing.TM2to1.tr_eval_dom {K : Type u_1} [DecidableEq K] {Γ : KType u_2} {Λ : Type u_3} [Inhabited Λ] {σ : Type u_4} [Inhabited σ] (M : ΛTuring.TM2.Stmt Γ Λ σ) (k : K) (L : List (Γ k)) :
                                                                                                                                                                                                    theorem Turing.TM2to1.tr_eval {K : Type u_1} [DecidableEq K] {Γ : KType u_2} {Λ : Type u_3} [Inhabited Λ] {σ : Type u_4} [Inhabited σ] (M : ΛTuring.TM2.Stmt Γ Λ σ) (k : K) (L : List (Γ k)) {L₁ : Turing.ListBlank Turing.TM2to1.Γ'} {L₂ : List (Γ k)} (H₁ : L₁ Turing.TM1.eval (Turing.TM2to1.tr M) (Turing.TM2to1.trInit k L)) (H₂ : L₂ Turing.TM2.eval M k L) :
                                                                                                                                                                                                    ∃ (S : (k : K) → List (Γ k)) (L' : Turing.ListBlank ((k : K) → Option (Γ k))), Turing.TM2to1.addBottom L' = L₁ (∀ (k : K), Turing.ListBlank.map (Turing.proj k) L' = Turing.ListBlank.mk (List.reverse (List.map some (S k)))) S k = L₂
                                                                                                                                                                                                    noncomputable def Turing.TM2to1.trSupp {K : Type u_1} {Γ : KType u_2} {Λ : Type u_3} {σ : Type u_4} (M : ΛTuring.TM2.Stmt Γ Λ σ) (S : Finset Λ) :
                                                                                                                                                                                                    Finset Turing.TM2to1.Λ'

                                                                                                                                                                                                    The support of a set of TM2 states in the TM2 emulator.

                                                                                                                                                                                                    Equations
                                                                                                                                                                                                    Instances For
                                                                                                                                                                                                      theorem Turing.TM2to1.tr_supports {K : Type u_1} [DecidableEq K] {Γ : KType u_2} {Λ : Type u_3} [Inhabited Λ] {σ : Type u_4} (M : ΛTuring.TM2.Stmt Γ Λ σ) {S : Finset Λ} (ss : Turing.TM2.Supports M S) :