Documentation

Mathlib.Computability.PostTuringMachine

Turing machines #

The files PostTuringMachine.lean and TuringMachine.lean define a sequence of simple machine languages, starting with Turing machines and working up to more complex languages based on Wang B-machines.

PostTuringMachine.lean covers the TM0 model and TM1 model; TuringMachine.lean adds the TM2 model.

Naming conventions #

Each model of computation in this file shares a naming convention for the elements of a model of computation. These are the parameters for the language:

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.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 : Reaches f a b) (hac : Reaches f a c) :
        Reaches f b c Reaches f c b
        theorem Turing.reaches₁_fwd {σ : Type u_1} {f : σOption σ} {a b c : σ} (h₁ : Reaches₁ f a c) (h₂ : b f a) :
        Reaches f b c
        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₁ : Reaches₀ f a b) (h₂ : 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₂ : Reaches₀ f b c) :
          theorem Turing.Reaches₀.tail {σ : Type u_1} {f : σOption σ} {a b c : σ} (h₁ : 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 : Reaches₁ f a b) :
          theorem Turing.Reaches.to₀ {σ : Type u_1} {f : σOption σ} {a b : σ} (h : Reaches f a b) :
          theorem Turing.Reaches₀.tail' {σ : Type u_1} {f : σOption σ} {a b c : σ} (h : Reaches₀ f a b) (h₂ : c f b) :
          def Turing.evalInduction {σ : Type u_2} {f : σOption σ} {b : σ} {C : σSort u_1} {a : σ} (h : b eval f a) (H : (a : σ) → b 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 eval f a Reaches f a b f b = none
            theorem Turing.eval_maximal₁ {σ : Type u_1} {f : σOption σ} {a b : σ} (h : b eval f a) (c : σ) :
            theorem Turing.eval_maximal {σ : Type u_1} {f : σOption σ} {a b : σ} (h : b eval f a) {c : σ} :
            Reaches f b c c = b
            theorem Turing.reaches_eval {σ : Type u_1} {f : σOption σ} {a b : σ} (ab : Reaches f a b) :
            eval f a = eval f 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
            Instances For
              theorem Turing.tr_reaches₁ {σ₁ : Type u_1} {σ₂ : Type u_2} {f₁ : σ₁Option σ₁} {f₂ : σ₂Option σ₂} {tr : σ₁σ₂Prop} (H : Respects f₁ f₂ tr) {a₁ : σ₁} {a₂ : σ₂} (aa : tr a₁ a₂) {b₁ : σ₁} (ab : Reaches₁ f₁ a₁ b₁) :
              ∃ (b₂ : σ₂), tr b₁ b₂ Reaches₁ f₂ a₂ b₂
              theorem Turing.tr_reaches {σ₁ : Type u_1} {σ₂ : Type u_2} {f₁ : σ₁Option σ₁} {f₂ : σ₂Option σ₂} {tr : σ₁σ₂Prop} (H : Respects f₁ f₂ tr) {a₁ : σ₁} {a₂ : σ₂} (aa : tr a₁ a₂) {b₁ : σ₁} (ab : Reaches f₁ a₁ b₁) :
              ∃ (b₂ : σ₂), tr b₁ b₂ Reaches f₂ a₂ b₂
              theorem Turing.tr_reaches_rev {σ₁ : Type u_1} {σ₂ : Type u_2} {f₁ : σ₁Option σ₁} {f₂ : σ₂Option σ₂} {tr : σ₁σ₂Prop} (H : Respects f₁ f₂ tr) {a₁ : σ₁} {a₂ : σ₂} (aa : tr a₁ a₂) {b₂ : σ₂} (ab : Reaches f₂ a₂ b₂) :
              ∃ (c₁ : σ₁) (c₂ : σ₂), Reaches f₂ b₂ c₂ tr c₁ c₂ Reaches f₁ a₁ c₁
              theorem Turing.tr_eval {σ₁ : Type u_1} {σ₂ : Type u_2} {f₁ : σ₁Option σ₁} {f₂ : σ₂Option σ₂} {tr : σ₁σ₂Prop} (H : Respects f₁ f₂ tr) {a₁ b₁ : σ₁} {a₂ : σ₂} (aa : tr a₁ a₂) (ab : b₁ eval f₁ a₁) :
              ∃ (b₂ : σ₂), tr b₁ b₂ b₂ eval f₂ a₂
              theorem Turing.tr_eval_rev {σ₁ : Type u_1} {σ₂ : Type u_2} {f₁ : σ₁Option σ₁} {f₂ : σ₂Option σ₂} {tr : σ₁σ₂Prop} (H : Respects f₁ f₂ tr) {a₁ : σ₁} {b₂ a₂ : σ₂} (aa : tr a₁ a₂) (ab : b₂ eval f₂ a₂) :
              ∃ (b₁ : σ₁), tr b₁ b₂ b₁ eval f₁ a₁
              theorem Turing.tr_eval_dom {σ₁ : Type u_1} {σ₂ : Type u_2} {f₁ : σ₁Option σ₁} {f₂ : σ₂Option σ₂} {tr : σ₁σ₂Prop} (H : Respects f₁ f₂ tr) {a₁ : σ₁} {a₂ : σ₂} (aa : tr a₁ a₂) :
              (eval f₂ a₂).Dom (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 σ₁} :
                FRespects f₂ tr a₂ b₁ FRespects f₂ tr b₂ b₁
                theorem Turing.fun_respects {σ₁ : Type u_1} {σ₂ : Type u_2} {f₁ : σ₁Option σ₁} {f₂ : σ₂Option σ₂} {tr : σ₁σ₂} :
                (Respects f₁ f₂ fun (a : σ₁) (b : σ₂) => tr a = b) ∀ ⦃a₁ : σ₁⦄, FRespects f₂ tr (tr a₁) (f₁ a₁)
                theorem Turing.tr_eval' {σ₁ σ₂ : Type u_1} (f₁ : σ₁Option σ₁) (f₂ : σ₂Option σ₂) (tr : σ₁σ₂) (H : Respects f₁ f₂ fun (a : σ₁) (b : σ₂) => tr a = b) (a₁ : σ₁) :
                eval f₂ (tr a₁) = tr <$> 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
                    structure Turing.TM0.Cfg (Γ : Type u_1) (Λ : Type u_2) [Inhabited Γ] :
                    Type (max u_1 u_2)

                    The configuration state of a Turing machine during operation consists of a label (machine state), and a tape. The tape is represented in the form (a, L, R), meaning the tape looks like L.rev ++ [a] ++ R with the machine currently reading the a. The lists are automatically extended with blanks as the machine moves around.

                    • q : Λ

                      The current machine state.

                    • Tape : Turing.Tape Γ

                      The current state of the tape: current symbol, left and right parts.

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

                      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} {Λ : Type u_2} [Inhabited Λ] [Inhabited Γ] (M : Machine Γ Λ) :
                        Cfg Γ Λ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} {Λ : Type u_2} [Inhabited Λ] [Inhabited Γ] (l : List Γ) :
                          Cfg Γ Λ

                          The initial configuration.

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

                                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 : PointedMap Γ Γ') (g : ΛΛ') :
                                  Cfg Γ Λ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 : Machine Γ Λ) (f₁ : PointedMap Γ Γ') (f₂ : PointedMap Γ' Γ) (g₁ : ΛΛ') (g₂ : Λ'Λ) :
                                    Machine Γ' Λ'

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

                                        • l : Option Λ

                                          The statement (if any) which is currently evaluated

                                        • var : σ

                                          The current value of the variable store

                                        • Tape : Turing.Tape Γ

                                          The current state of the tape

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

                                          The semantics of TM1 evaluation.

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

                                            The state transition function.

                                            Equations
                                            Instances For
                                              def Turing.TM1.SupportsStmt {Γ : Type u_1} {Λ : Type u_2} {σ : Type u_3} (S : Finset Λ) :
                                              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 : Stmt Γ Λ σ} :
                                                theorem Turing.TM1.stmts₁_trans {Γ : Type u_1} {Λ : Type u_2} {σ : Type u_3} {q₁ q₂ : Stmt Γ Λ σ} :
                                                q₁ stmts₁ q₂stmts₁ q₁ stmts₁ q₂
                                                theorem Turing.TM1.stmts₁_supportsStmt_mono {Γ : Type u_1} {Λ : Type u_2} {σ : Type u_3} {S : Finset Λ} {q₁ q₂ : Stmt Γ Λ σ} (h : q₁ stmts₁ q₂) (hs : SupportsStmt S q₂) :
                                                noncomputable def Turing.TM1.stmts {Γ : Type u_1} {Λ : Type u_2} {σ : Type u_3} (M : ΛStmt Γ Λ σ) (S : Finset Λ) :
                                                Finset (Option (Stmt Γ Λ σ))

                                                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 : ΛStmt Γ Λ σ} {S : Finset Λ} {q₁ q₂ : Stmt Γ Λ σ} (h₁ : q₁ stmts₁ q₂) :
                                                  some q₂ stmts M Ssome q₁ stmts M S
                                                  def Turing.TM1.Supports {Γ : Type u_1} {Λ : Type u_2} {σ : Type u_3} [Inhabited Λ] (M : Λ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 : ΛStmt Γ Λ σ} {S : Finset Λ} {q : Stmt Γ Λ σ} (ss : Supports M S) :
                                                    some q stmts M SSupportsStmt S q
                                                    theorem Turing.TM1.step_supports {Γ : Type u_1} {Λ : Type u_2} {σ : Type u_3} [Inhabited Λ] [Inhabited Γ] (M : ΛStmt Γ Λ σ) {S : Finset Λ} (ss : Supports M S) {c c' : Cfg Γ Λ σ} :
                                                    def Turing.TM1.init {Γ : Type u_1} {Λ : Type u_2} {σ : Type u_3} [Inhabited Λ] [Inhabited Γ] [Inhabited σ] (l : List Γ) :
                                                    Cfg Γ Λ σ

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

                                                          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 : ΛTM1.Stmt Γ Λ σ) :

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

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

                                                              Translate configurations from TM1 to TM0.

                                                              Equations
                                                              Instances For
                                                                theorem Turing.TM1to0.tr_respects {Γ : Type u_1} {Λ : Type u_2} [Inhabited Λ] {σ : Type u_3} [Inhabited σ] (M : ΛTM1.Stmt Γ Λ σ) [Inhabited Γ] :
                                                                Respects (TM1.step M) (TM0.step (tr M)) fun (c₁ : TM1.Cfg Γ Λ σ) (c₂ : TM0.Cfg Γ (Λ' M)) => trCfg M c₁ = c₂
                                                                theorem Turing.TM1to0.tr_eval {Γ : Type u_1} {Λ : Type u_2} [Inhabited Λ] {σ : Type u_3} [Inhabited σ] (M : ΛTM1.Stmt Γ Λ σ) [Inhabited Γ] (l : List Γ) :
                                                                TM0.eval (tr M) l = TM1.eval M l
                                                                noncomputable def Turing.TM1to0.trStmts {Γ : Type u_1} {Λ : Type u_2} {σ : Type u_3} (M : Λ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 : ΛTM1.Stmt Γ Λ σ) [Fintype σ] {S : Finset Λ} (ss : TM1.Supports M S) :
                                                                  TM0.Supports (tr M) (trStmts 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 : ΓList.Vector Bool n) (dec : List.Vector Bool nΓ), enc default = List.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.

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

                                                                    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 : Dir) (q : TM1.Stmt Bool (Λ' Γ Λ σ) σ) :
                                                                      TM1.Stmt Bool (Λ' Γ Λ σ) σ

                                                                      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 : List.Vector Bool nΓ) (f : ΓTM1.Stmt Bool (Λ' Γ Λ σ) σ) :
                                                                        TM1.Stmt Bool (Λ' Γ Λ σ) σ

                                                                        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 BoolTM1.Stmt Bool (Λ' Γ Λ σ) σTM1.Stmt Bool (Λ' Γ Λ σ) σ

                                                                          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 : List.Vector Bool nΓ) :
                                                                            TM1.Stmt Γ Λ σTM1.Stmt Bool (Λ' Γ Λ σ) σ

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

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

                                                                              Equations
                                                                              Instances For
                                                                                def Turing.TM1to1.trTape {Γ : Type u_1} {n : } {enc : ΓList.Vector Bool n} [Inhabited Γ] (enc0 : enc default = List.Vector.replicate n false) (T : Tape Γ) :

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

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

                                                                                  The top level program.

                                                                                  Equations
                                                                                  Instances For
                                                                                    def Turing.TM1to1.trCfg {Γ : Type u_1} {Λ : Type u_2} {σ : Type u_3} {n : } (enc : ΓList.Vector Bool n) [Inhabited Γ] (enc0 : enc default = List.Vector.replicate n false) :
                                                                                    TM1.Cfg Γ Λ σTM1.Cfg Bool (Λ' Γ Λ σ) σ

                                                                                    The machine configuration translation.

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

                                                                                      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} {σ : Type u_3} {n : } (enc : ΓList.Vector Bool n) (dec : List.Vector Bool nΓ) (M : ΛTM1.Stmt Γ Λ σ) [Fintype Γ] [Inhabited Λ] {S : Finset Λ} (ss : TM1.Supports M S) :
                                                                                        TM1.Supports (tr enc dec M) (trSupp 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.

                                                                                        Instances For
                                                                                          def Turing.TM0to1.tr {Γ : Type u_1} {Λ : Type u_2} [Inhabited Λ] (M : TM0.Machine Γ Λ) :
                                                                                          Λ' Γ ΛTM1.Stmt Γ (Λ' Γ Λ) Unit

                                                                                          The program.

                                                                                          Equations
                                                                                          Instances For
                                                                                            def Turing.TM0to1.trCfg {Γ : Type u_1} [Inhabited Γ] {Λ : Type u_2} [Inhabited Λ] (M : TM0.Machine Γ Λ) :
                                                                                            TM0.Cfg Γ ΛTM1.Cfg Γ (Λ' Γ Λ) Unit

                                                                                            The configuration translation.

                                                                                            Equations
                                                                                            Instances For
                                                                                              theorem Turing.TM0to1.tr_respects {Γ : Type u_1} [Inhabited Γ] {Λ : Type u_2} [Inhabited Λ] (M : TM0.Machine Γ Λ) :
                                                                                              Respects (TM0.step M) (TM1.step (tr M)) fun (a : TM0.Cfg Γ Λ) (b : TM1.Cfg Γ (Λ' Γ Λ) Unit) => trCfg M a = b