Documentation

Mathlib.Computability.TMToPartrec

Modelling partial recursive functions using Turing machines #

The files TMConfig and TMToPartrec define a simplified basis for partial recursive functions, and a Turing.TM2 model Turing machine for evaluating these functions. This amounts to a constructive proof that every Partrec function can be evaluated by a Turing machine.

Main definitions #

Simulating sequentialized partial recursive functions in TM2 #

At this point we have a sequential model of partial recursive functions: the Cfg type and step : Cfg → Option Cfg function from TMConfig.lean. The key feature of this model is that it does a finite amount of computation (in fact, an amount which is statically bounded by the size of the program) between each step, and no individual step can diverge (unlike the compositional semantics, where every sub-part of the computation is potentially divergent). So we can utilize the same techniques as in the other TM simulations in Computability.TuringMachine to prove that each step corresponds to a finite number of steps in a lower level model. (We don't prove it here, but in anticipation of the complexity class P, the simulation is actually polynomial-time as well.)

The target model is Turing.TM2, which has a fixed finite set of stacks, a bit of local storage, with programs selected from a potentially infinite (but finitely accessible) set of program positions, or labels Λ, each of which executes a finite sequence of basic stack commands.

For this program we will need four stacks, each on an alphabet Γ' like so:

inductive Γ'  | consₗ | cons | bit0 | bit1

We represent a number as a bit sequence, lists of numbers by putting cons after each element, and lists of lists of natural numbers by putting consₗ after each list. For example:

0 ~> []
1 ~> [bit1]
6 ~> [bit0, bit1, bit1]
[1, 2] ~> [bit1, cons, bit0, bit1, cons]
[[], [1, 2]] ~> [consₗ, bit1, cons, bit0, bit1, cons, consₗ]

The four stacks are main, rev, aux, stack. In normal mode, main contains the input to the current program (a List) and stack contains data (a List (List ℕ)) associated to the current continuation, and in ret mode main contains the value that is being passed to the continuation and stack contains the data for the continuation. The rev and aux stacks are usually empty; rev is used to store reversed data when e.g. moving a value from one stack to another, while aux is used as a temporary for a main/stack swap that happens during cons₁ evaluation.

The only local store we need is Option Γ', which stores the result of the last pop operation. (Most of our working data are natural numbers, which are too large to fit in the local store.)

The continuations from the previous section are data-carrying, containing all the values that have been computed and are awaiting other arguments. In order to have only a finite number of continuations appear in the program so that they can be used in machine states, we separate the data part (anything with type List) from the Cont type, producing a Cont' type that lacks this information. The data is kept on the stack stack.

Because we want to have subroutines for e.g. moving an entire stack to another place, we use an infinite inductive type Λ' so that we can execute a program and then return to do something else without having to define too many different kinds of intermediate states. (We must nevertheless prove that only finitely many labels are accessible.) The labels are:

In addition to these basic states, we define some additional subroutines that are used in the above:

The main theorem of this section is tr_eval, which asserts that for each that for each code c, the state init c v steps to halt v' in finitely many steps if and only if Code.eval c v = some v'.

The alphabet for the stacks in the program. bit0 and bit1 are used to represent values as lists of binary digits, cons is used to separate List values, and consₗ is used to separate List (List ℕ) values. See the section documentation.

Instances For
    Equations
    • One or more equations did not get rendered due to their size.

    The four stacks used by the program. main is used to store the input value in trNormal mode and the output value in Λ'.ret mode, while stack is used to keep all the data for the continuations. rev is used to store reversed lists when transferring values between stacks, and aux is only used once in cons₁. See the section documentation.

    Instances For

      Continuations as in ToPartrec.Cont but with the data removed. This is done because we want the set of all continuations in the program to be finite (so that it can ultimately be encoded into the finite state machine of a Turing machine), but a continuation can handle a potentially infinite number of data values during execution.

      Instances For

        The set of program positions. We make extensive use of inductive types here to let us describe "subroutines"; for example clear p k q is a program that clears stack k, then does q where q is another label. In order to prevent this from resulting in an infinite number of distinct accessible states, we are careful to be non-recursive (although loops are okay). See the section documentation for a description of all the programs.

        Instances For
          Equations
          • One or more equations did not get rendered due to their size.

          A predicate that detects the end of a natural number, either Γ'.cons or Γ'.consₗ (or implicitly the end of the list), for use in predicate-taking functions like move and clear.

          Equations
          Instances For

            Pop a value from the stack and place the result in local store.

            Equations
            Instances For

              Peek a value from the stack and place the result in local store.

              Equations
              Instances For

                Push the value in the local store to the given stack.

                Equations
                Instances For
                  def Turing.PartrecToTM2.moveExcl (p : Γ'Bool) (k₁ k₂ : K') (q : Λ') :

                  Move elements from k₁ to k₂ while p holds, with the last element being left on k₁.

                  Equations
                  Instances For
                    def Turing.PartrecToTM2.move₂ (p : Γ'Bool) (k₁ k₂ : K') (q : Λ') :

                    Move elements from k₁ to k₂ without reversion, by performing a double move via the rev stack.

                    Equations
                    • One or more equations did not get rendered due to their size.
                    Instances For

                      Assuming trList v is on the front of stack k, remove it, and push v.headI onto main. See the section documentation.

                      Equations
                      • One or more equations did not get rendered due to their size.
                      Instances For
                        @[simp]
                        theorem Turing.PartrecToTM2.tr_move (p : Γ'Bool) (k₁ k₂ : K') (q : Λ') :
                        tr (Λ'.move p k₁ k₂ q) = pop' k₁ (TM2.Stmt.branch (fun (s : Option Γ') => s.elim true p) (TM2.Stmt.goto fun (x : Option Γ') => q) (push' k₂ (TM2.Stmt.goto fun (x : Option Γ') => Λ'.move p k₁ k₂ q)))
                        @[simp]
                        theorem Turing.PartrecToTM2.tr_push (k : K') (f : Option Γ'Option Γ') (q : Λ') :
                        tr (Λ'.push k f q) = TM2.Stmt.branch (fun (s : Option Γ') => (f s).isSome) (TM2.Stmt.push k (fun (s : Option Γ') => (f s).iget) (TM2.Stmt.goto fun (x : Option Γ') => q)) (TM2.Stmt.goto fun (x : Option Γ') => q)
                        @[simp]
                        theorem Turing.PartrecToTM2.tr_clear (p : Γ'Bool) (k : K') (q : Λ') :
                        tr (Λ'.clear p k q) = pop' k (TM2.Stmt.branch (fun (s : Option Γ') => s.elim true p) (TM2.Stmt.goto fun (x : Option Γ') => q) (TM2.Stmt.goto fun (x : Option Γ') => Λ'.clear p k q))
                        @[simp]
                        theorem Turing.PartrecToTM2.tr_pred (q₁ q₂ : Λ') :
                        tr (q₁.pred q₂) = pop' K'.main (TM2.Stmt.branch (fun (s : Option Γ') => decide (s = some Γ'.bit0)) (TM2.Stmt.push K'.rev (fun (x : Option Γ') => Γ'.bit1) (TM2.Stmt.goto fun (x : Option Γ') => q₁.pred q₂)) (TM2.Stmt.branch (fun (s : Option Γ') => natEnd s.iget) (TM2.Stmt.goto fun (x : Option Γ') => q₁) (peek' K'.main (TM2.Stmt.branch (fun (s : Option Γ') => natEnd s.iget) (TM2.Stmt.goto fun (x : Option Γ') => unrev q₂) (TM2.Stmt.push K'.rev (fun (x : Option Γ') => Γ'.bit0) (TM2.Stmt.goto fun (x : Option Γ') => unrev q₂))))))

                        We use PosNum to define the translation of binary natural numbers. A natural number is represented as a little-endian list of bit0 and bit1 elements:

                        1 = [bit1]
                        2 = [bit0, bit1]
                        3 = [bit1, bit1]
                        4 = [bit0, bit0, bit1]
                        

                        In particular, this representation guarantees no trailing bit0's at the end of the list.

                        Equations
                        Instances For

                          We use Num to define the translation of binary natural numbers. Positive numbers are translated using trPosNum, and trNum 0 = []. So there are never any trailing bit0's in a translated Num.

                          0 = []
                          1 = [bit1]
                          2 = [bit0, bit1]
                          3 = [bit1, bit1]
                          4 = [bit0, bit0, bit1]
                          
                          Equations
                          Instances For

                            Because we use binary encoding, we define trNat in terms of trNum, using Num, which are binary natural numbers. (We could also use Nat.binaryRecOn, but Num and PosNum make for easy inductions.)

                            Equations
                            Instances For

                              Lists are translated with a cons after each encoded number. For example:

                              [] = []
                              [0] = [cons]
                              [1] = [bit1, cons]
                              [6, 0] = [bit0, bit1, bit1, cons, cons]
                              
                              Equations
                              Instances For

                                Lists of lists are translated with a consₗ after each encoded list. For example:

                                [] = []
                                [[]] = [consₗ]
                                [[], []] = [consₗ, consₗ]
                                [[0]] = [cons, consₗ]
                                [[1, 2], [0]] = [bit1, cons, bit0, bit1, cons, consₗ, cons, consₗ]
                                
                                Equations
                                Instances For

                                  The data part of a continuation is a list of lists, which is encoded on the stack stack using trLList.

                                  Equations
                                  Instances For

                                    This is the nondependent eliminator for K', but we use it specifically here in order to represent the stack data as four lists rather than as a function K'List Γ', because this makes rewrites easier. The theorems K'.elim_update_main et. al. show how such a function is updated after an update to one of the components.

                                    Equations
                                    Instances For
                                      theorem Turing.PartrecToTM2.K'.elim_main (a b c d : List Γ') :
                                      elim a b c d main = a
                                      theorem Turing.PartrecToTM2.K'.elim_rev (a b c d : List Γ') :
                                      elim a b c d rev = b
                                      theorem Turing.PartrecToTM2.K'.elim_aux (a b c d : List Γ') :
                                      elim a b c d aux = c
                                      @[simp]
                                      theorem Turing.PartrecToTM2.K'.elim_update_main {a b c d a' : List Γ'} :
                                      Function.update (elim a b c d) main a' = elim a' b c d
                                      @[simp]
                                      theorem Turing.PartrecToTM2.K'.elim_update_rev {a b c d b' : List Γ'} :
                                      Function.update (elim a b c d) rev b' = elim a b' c d
                                      @[simp]
                                      theorem Turing.PartrecToTM2.K'.elim_update_aux {a b c d c' : List Γ'} :
                                      Function.update (elim a b c d) aux c' = elim a b c' d
                                      @[simp]
                                      theorem Turing.PartrecToTM2.K'.elim_update_stack {a b c d d' : List Γ'} :
                                      Function.update (elim a b c d) stack d' = elim a b c d'

                                      The halting state corresponding to a List output value.

                                      Equations
                                      Instances For

                                        The Cfg states map to Cfg' states almost one to one, except that in normal operation the local store contains an arbitrary garbage value. To make the final theorem cleaner we explicitly clear it in the halt state so that there is exactly one configuration corresponding to output v.

                                        Equations
                                        Instances For
                                          def Turing.PartrecToTM2.splitAtPred {α : Type u_1} (p : αBool) :
                                          List αList α × Option α × List α

                                          This could be a general list definition, but it is also somewhat specialized to this application. splitAtPred p L will search L for the first element satisfying p. If it is found, say L = l₁ ++ a :: l₂ where a satisfies p but l₁ does not, then it returns (l₁, some a, l₂). Otherwise, if there is no such element, it returns (L, none, []).

                                          Equations
                                          Instances For
                                            theorem Turing.PartrecToTM2.splitAtPred_eq {α : Type u_1} (p : αBool) (L l₁ : List α) (o : Option α) (l₂ : List α) :
                                            (∀ xl₁, p x = false)Option.elim' (L = l₁ l₂ = []) (fun (a : α) => p a = true L = l₁ ++ a :: l₂) osplitAtPred p L = (l₁, o, l₂)
                                            theorem Turing.PartrecToTM2.splitAtPred_false {α : Type u_1} (L : List α) :
                                            splitAtPred (fun (x : α) => false) L = (L, none, [])
                                            theorem Turing.PartrecToTM2.move_ok {p : Γ'Bool} {k₁ k₂ : K'} {q : Λ'} {s : Option Γ'} {L₁ : List Γ'} {o : Option Γ'} {L₂ : List Γ'} {S : K'List Γ'} (h₁ : k₁ k₂) (e : splitAtPred p (S k₁) = (L₁, o, L₂)) :
                                            Reaches₁ (TM2.step tr) { l := some (Λ'.move p k₁ k₂ q), var := s, stk := S } { l := some q, var := o, stk := Function.update (Function.update S k₁ L₂) k₂ (L₁.reverseAux (S k₂)) }
                                            theorem Turing.PartrecToTM2.unrev_ok {q : Λ'} {s : Option Γ'} {S : K'List Γ'} :
                                            Reaches₁ (TM2.step tr) { l := some (unrev q), var := s, stk := S } { l := some q, var := none, stk := Function.update (Function.update S K'.rev []) K'.main ((S K'.rev).reverseAux (S K'.main)) }
                                            theorem Turing.PartrecToTM2.move₂_ok {p : Γ'Bool} {k₁ k₂ : K'} {q : Λ'} {s : Option Γ'} {L₁ : List Γ'} {o : Option Γ'} {L₂ : List Γ'} {S : K'List Γ'} (h₁ : k₁ K'.rev k₂ K'.rev k₁ k₂) (h₂ : S K'.rev = []) (e : splitAtPred p (S k₁) = (L₁, o, L₂)) :
                                            Reaches₁ (TM2.step tr) { l := some (move₂ p k₁ k₂ q), var := s, stk := S } { l := some q, var := none, stk := Function.update (Function.update S k₁ (o.elim id List.cons L₂)) k₂ (L₁ ++ S k₂) }
                                            theorem Turing.PartrecToTM2.clear_ok {p : Γ'Bool} {k : K'} {q : Λ'} {s : Option Γ'} {L₁ : List Γ'} {o : Option Γ'} {L₂ : List Γ'} {S : K'List Γ'} (e : splitAtPred p (S k) = (L₁, o, L₂)) :
                                            Reaches₁ (TM2.step tr) { l := some (Λ'.clear p k q), var := s, stk := S } { l := some q, var := o, stk := Function.update S k L₂ }
                                            theorem Turing.PartrecToTM2.copy_ok (q : Λ') (s : Option Γ') (a b c d : List Γ') :
                                            Reaches₁ (TM2.step tr) { l := some q.copy, var := s, stk := K'.elim a b c d } { l := some q, var := none, stk := K'.elim (b.reverseAux a) [] c (b.reverseAux d) }
                                            theorem Turing.PartrecToTM2.head_main_ok {q : Λ'} {s : Option Γ'} {L : List } {c d : List Γ'} :
                                            Reaches₁ (TM2.step tr) { l := some (head K'.main q), var := s, stk := K'.elim (trList L) [] c d } { l := some q, var := none, stk := K'.elim (trList [L.headI]) [] c d }
                                            theorem Turing.PartrecToTM2.head_stack_ok {q : Λ'} {s : Option Γ'} {L₁ L₂ : List } {L₃ : List Γ'} :
                                            Reaches₁ (TM2.step tr) { l := some (head K'.stack q), var := s, stk := K'.elim (trList L₁) [] [] (trList L₂ ++ Γ'.consₗ :: L₃) } { l := some q, var := none, stk := K'.elim (trList (L₂.headI :: L₁)) [] [] L₃ }
                                            theorem Turing.PartrecToTM2.succ_ok {q : Λ'} {s : Option Γ'} {n : } {c d : List Γ'} :
                                            Reaches₁ (TM2.step tr) { l := some q.succ, var := s, stk := K'.elim (trList [n]) [] c d } { l := some q, var := none, stk := K'.elim (trList [n.succ]) [] c d }
                                            theorem Turing.PartrecToTM2.pred_ok (q₁ q₂ : Λ') (s : Option Γ') (v : List ) (c d : List Γ') :
                                            ∃ (s' : Option Γ'), Reaches₁ (TM2.step tr) { l := some (q₁.pred q₂), var := s, stk := K'.elim (trList v) [] c d } (Nat.rec { l := some q₁, var := s', stk := K'.elim (trList v.tail) [] c d } (fun (n : ) (x : TM2.Cfg (fun (x : K') => Γ') Λ' (Option Γ')) => { l := some q₂, var := s', stk := K'.elim (trList (n :: v.tail)) [] c d }) v.headI)
                                            theorem Turing.PartrecToTM2.trNormal_respects (c : ToPartrec.Code) (k : ToPartrec.Cont) (v : List ) (s : Option Γ') :
                                            ∃ (b₂ : Cfg'), TrCfg (ToPartrec.stepNormal c k v) b₂ Reaches₁ (TM2.step tr) { l := some (trNormal c (trCont k)), var := s, stk := K'.elim (trList v) [] [] (trContStack k) } b₂
                                            theorem Turing.PartrecToTM2.tr_ret_respects (k : ToPartrec.Cont) (v : List ) (s : Option Γ') :
                                            ∃ (b₂ : Cfg'), TrCfg (ToPartrec.stepRet k v) b₂ Reaches₁ (TM2.step tr) { l := some (Λ'.ret (trCont k)), var := s, stk := K'.elim (trList v) [] [] (trContStack k) } b₂

                                            The initial state, evaluating function c on input v.

                                            Equations
                                            • One or more equations did not get rendered due to their size.
                                            Instances For

                                              The (finite!) set of machine states visited during the course of evaluation of c, including the state ret k but not any states after that (that is, the states visited while evaluating k).

                                              Equations
                                              Instances For

                                                The (finite!) set of machine states visited during the course of evaluation of c in continuation k. This is actually closed under forward simulation (see tr_supports), and the existence of this set means that the machine constructed in this section is in fact a proper Turing machine, with a finite set of states.

                                                Equations
                                                Instances For

                                                  A shorthand for the predicate that we are proving in the main theorems trStmts₁_supports, codeSupp'_supports, contSupp_supports, codeSupp_supports. The set S is fixed throughout the proof, and denotes the full set of states in the machine, while K is a subset that we are currently proving a property about. The predicate asserts that every state in K is closed in S under forward simulation, i.e. stepping forward through evaluation starting from any state in K stays entirely within S.

                                                  Equations
                                                  Instances For
                                                    theorem Turing.PartrecToTM2.supports_union {K₁ K₂ S : Finset Λ'} :
                                                    Supports (K₁ K₂) S Supports K₁ S Supports K₂ S
                                                    theorem Turing.PartrecToTM2.trStmts₁_supports' {S : Finset Λ'} {q : Λ'} {K : Finset Λ'} (H₁ : Λ'.Supports S q) (H₂ : trStmts₁ q K S) (H₃ : K SSupports K S) :

                                                    The set codeSupp c k is a finite set that witnesses the effective finiteness of the tr Turing machine. Starting from the initial state trNormal c k, forward simulation uses only states in codeSupp c k, so this is a finite state machine. Even though the underlying type of state labels Λ' is infinite, for a given partial recursive function c and continuation k, only finitely many states are accessed, corresponding roughly to subterms of c.