Computable functions #
This file contains the definition of a Turing machine with some finiteness conditions
(bundling the definition of TM2 in TuringMachine.lean
), a definition of when a TM gives a certain
output (in a certain time), and the definition of computability (in polynomial time or
any time function) of a function between two types that have an encoding (as in Encoding.lean
).
Main theorems #
idComputableInPolyTime
: a TM + a proof it computes the identity on a type in polytime.idComputable
: a TM + a proof it computes the identity on a type.
Implementation notes #
To count the execution time of a Turing machine, we have decided to count the number of times the
step
function is used. Each step executes a statement (of type Stmt
); this is a function, and
generally contains multiple "fundamental" steps (pushing, popping, and so on).
However, as functions only contain a finite number of executions and each one is executed at most
once, this execution time is up to multiplication by a constant the amount of fundamental steps.
A bundled TM2 (an equivalent of the classical Turing machine, defined starting from
the namespace Turing.TM2
in TuringMachine.lean
), with an input and output stack,
a main function, an initial state and some finiteness guarantees.
- K : Type
index type of stacks
- kDecidableEq : DecidableEq self.K
A TM2 machine has finitely many stacks.
- k₀ : self.K
input resp. output stack
- k₁ : self.K
input resp. output stack
type of stack elements
- Λ : Type
type of function labels
- main : self.Λ
a main function: the initial function that is executed, given by its label
A TM2 machine has finitely many function labels.
- σ : Type
type of states of the machine
- initialState : self.σ
the initial state of the machine
a TM2 machine has finitely many internal states.
Each internal stack is finite.
the program itself, i.e. one function for every function label
Instances For
Equations
- tm.decidableEqK = tm.kDecidableEq
Equations
- tm.inhabitedσ = { default := tm.initialState }
The type of statements (functions) corresponding to this TM.
Instances For
Equations
- tm.inhabitedStmt = inferInstanceAs (Inhabited (Turing.TM2.Stmt tm.Γ tm.Λ tm.σ))
The type of configurations (functions) corresponding to this TM.
Instances For
Equations
- tm.inhabitedCfg = Turing.TM2.Cfg.inhabited tm.Γ tm.Λ tm.σ
A "proof" of the fact that f
eventually reaches b
when repeatedly evaluated on a
,
remembering the number of steps it takes.
- steps : ℕ
number of steps taken
Instances For
A "proof" of the fact that f
eventually reaches b
in at most m
steps when repeatedly
evaluated on a
, remembering the number of steps it takes.
Instances For
Reflexivity of EvalsTo
in 0 steps.
Equations
- Turing.EvalsTo.refl f a = { steps := 0, evals_in_steps := ⋯ }
Instances For
Transitivity of EvalsTo
in the sum of the numbers of steps.
Equations
- Turing.EvalsTo.trans f a b c h₁ h₂ = { steps := h₂.steps + h₁.steps, evals_in_steps := ⋯ }
Instances For
Reflexivity of EvalsToInTime
in 0 steps.
Equations
- Turing.EvalsToInTime.refl f a = { toEvalsTo := Turing.EvalsTo.refl f a, steps_le_m := Turing.EvalsToInTime.refl.proof_1 }
Instances For
Transitivity of EvalsToInTime
in the sum of the numbers of steps.
Equations
- Turing.EvalsToInTime.trans f m₁ m₂ a b c h₁ h₂ = { toEvalsTo := Turing.EvalsTo.trans f a b c h₁.toEvalsTo h₂.toEvalsTo, steps_le_m := ⋯ }
Instances For
A proof of tm outputting l' when given l.
Equations
- Turing.TM2Outputs tm l l' = Turing.EvalsTo tm.step (Turing.initList tm l) (Option.map (Turing.haltList tm) l')
Instances For
A proof of tm outputting l' when given l in at most m steps.
Equations
- Turing.TM2OutputsInTime tm l l' m = Turing.EvalsToInTime tm.step (Turing.initList tm l) (Option.map (Turing.haltList tm) l') m
Instances For
The forgetful map, forgetting the upper bound on the number of steps.
Equations
- h.toTM2Outputs = h.toEvalsTo
Instances For
A (bundled TM2) Turing machine
with input alphabet equivalent to Γ₀
and output alphabet equivalent to Γ₁
.
- tm : FinTM2
the underlying bundled TM2
the input alphabet is equivalent to
Γ₀
the output alphabet is equivalent to
Γ₁
Instances For
A Turing machine + a proof it outputs f
.
- outputsFun (a : α) : TM2Outputs self.tm (List.map self.inputAlphabet.invFun (ea.encode a)) (some (List.map self.outputAlphabet.invFun (eb.encode (f a))))
a proof this machine outputs
f
Instances For
A Turing machine + a time function +
a proof it outputs f
in at most time(input.length)
steps.
a time function
- outputsFun (a : α) : TM2OutputsInTime self.tm (List.map self.inputAlphabet.invFun (ea.encode a)) (some (List.map self.outputAlphabet.invFun (eb.encode (f a)))) (self.time (ea.encode a).length)
proof this machine outputs
f
in at mosttime(input.length)
steps
Instances For
A Turing machine + a polynomial time function +
a proof it outputs f
in at most time(input.length)
steps.
- time : Polynomial ℕ
a polynomial time function
- outputsFun (a : α) : TM2OutputsInTime self.tm (List.map self.inputAlphabet.invFun (ea.encode a)) (some (List.map self.outputAlphabet.invFun (eb.encode (f a)))) (Polynomial.eval (ea.encode a).length self.time)
proof that this machine outputs
f
in at mosttime(input.length)
steps
Instances For
A forgetful map, forgetting the time bound on the number of steps.
Equations
- h.toTM2Computable = { toTM2ComputableAux := h.toTM2ComputableAux, outputsFun := fun (a : α) => (h.outputsFun a).toTM2Outputs }
Instances For
A forgetful map, forgetting that the time function is polynomial.
Equations
- h.toTM2ComputableInTime = { toTM2ComputableAux := h.toTM2ComputableAux, time := fun (n : ℕ) => Polynomial.eval n h.time, outputsFun := h.outputsFun }
Instances For
A Turing machine computing the identity on α.
Equations
- Turing.idComputer ea = Turing.FinTM2.mk PUnit.unit PUnit.unit (fun (x : Unit) => ea.Γ) Unit PUnit.unit Unit PUnit.unit fun (x : Unit) => Turing.TM2.Stmt.halt
Instances For
Equations
- Turing.inhabitedFinTM2 = { default := Turing.idComputer default }
A proof that the identity map on α is computable in polytime.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
Equations
Equations
- Turing.inhabitedTM2Outputs = { default := default.toTM2Outputs }
Equations
- Turing.inhabitedEvalsToInTime = { default := Turing.EvalsToInTime.refl (fun (x : Unit) => some PUnit.unit) PUnit.unit }
Equations
- Turing.inhabitedTM2EvalsTo = { default := Turing.EvalsTo.refl (fun (x : Unit) => some PUnit.unit) PUnit.unit }
A proof that the identity map on α is computable in time.
Equations
Instances For
A proof that the identity map on α is computable.
Equations
Instances For
Equations
- Turing.inhabitedTM2ComputableAux = { default := default.toTM2ComputableAux }