Documentation

Mathlib.Computability.NFA

Nondeterministic Finite Automata #

A Nondeterministic Finite Automaton (NFA) is a state machine which decides membership in a particular Language, by following every possible path that describes an input string.

We show that DFAs and NFAs can decide the same languages, by constructing an equivalent DFA for every NFA, and vice versa.

As constructing a DFA from an NFA uses an exponential number of states, we re-prove the pumping lemma instead of lifting DFA.pumping_lemma, in order to obtain the optimal bound on the minimal length of the string.

Like DFA, this definition allows for automata with infinite states; a Fintype instance must be supplied for true NFAs.

Main definitions #

Main theorems #

structure NFA (α : Type u) (σ : Type v) :
Type (max u v)

An NFA is a set of states (σ), a transition function from state to state labelled by the alphabet (step), a set of starting states (start) and a set of acceptance states (accept). Note the transition function sends a state to a Set of states. These are the states that it may be sent to.

  • step : σαSet σ

    The NFA's transition function

  • start : Set σ

    Set of starting states

  • accept : Set σ

    Set of accepting states

Instances For
    instance NFA.instInhabited {α : Type u} {σ : Type v} :
    Inhabited (NFA α σ)
    Equations
    def NFA.stepSet {α : Type u} {σ : Type v} (M : NFA α σ) (S : Set σ) (a : α) :
    Set σ

    M.stepSet S a is the union of M.step s a for all s ∈ S.

    Equations
    Instances For
      theorem NFA.mem_stepSet {α : Type u} {σ : Type v} {M : NFA α σ} {s : σ} {S : Set σ} {a : α} :
      s M.stepSet S a tS, s M.step t a
      @[simp]
      theorem NFA.stepSet_empty {α : Type u} {σ : Type v} (M : NFA α σ) (a : α) :
      @[simp]
      theorem NFA.stepSet_singleton {α : Type u} {σ : Type v} (M : NFA α σ) (s : σ) (a : α) :
      M.stepSet {s} a = M.step s a
      def NFA.evalFrom {α : Type u} {σ : Type v} (M : NFA α σ) (S : Set σ) :
      List αSet σ

      M.evalFrom S x computes all possible paths through M with input x starting at an element of S.

      Equations
      Instances For
        @[simp]
        theorem NFA.evalFrom_nil {α : Type u} {σ : Type v} (M : NFA α σ) (S : Set σ) :
        M.evalFrom S [] = S
        @[simp]
        theorem NFA.evalFrom_singleton {α : Type u} {σ : Type v} (M : NFA α σ) (S : Set σ) (a : α) :
        M.evalFrom S [a] = M.stepSet S a
        @[simp]
        theorem NFA.evalFrom_cons {α : Type u} {σ : Type v} (M : NFA α σ) (S : Set σ) (a : α) (x : List α) :
        M.evalFrom S (a :: x) = M.evalFrom (M.stepSet S a) x
        @[simp]
        theorem NFA.evalFrom_append_singleton {α : Type u} {σ : Type v} {M : NFA α σ} (S : Set σ) (x : List α) (a : α) :
        M.evalFrom S (x ++ [a]) = M.stepSet (M.evalFrom S x) a
        @[simp]
        theorem NFA.evalFrom_biUnion {α : Type u} {σ : Type v} (M : NFA α σ) {ι : Type u_1} (t : Set ι) (f : ιSet σ) (x : List α) :
        M.evalFrom (⋃ it, f i) x = it, M.evalFrom (f i) x
        theorem NFA.evalFrom_eq_biUnion_singleton {α : Type u} {σ : Type v} (M : NFA α σ) (S : Set σ) (x : List α) :
        M.evalFrom S x = sS, M.evalFrom {s} x
        theorem NFA.mem_evalFrom_iff_exists {α : Type u} {σ : Type v} {M : NFA α σ} {s : σ} {S : Set σ} {x : List α} :
        s M.evalFrom S x tS, s M.evalFrom {t} x
        def NFA.eval {α : Type u} {σ : Type v} (M : NFA α σ) :
        List αSet σ

        M.eval x computes all possible paths though M with input x starting at an element of M.start.

        Equations
        Instances For
          @[simp]
          theorem NFA.eval_nil {α : Type u} {σ : Type v} (M : NFA α σ) :
          @[simp]
          theorem NFA.eval_singleton {α : Type u} {σ : Type v} (M : NFA α σ) (a : α) :
          M.eval [a] = M.stepSet M.start a
          @[simp]
          theorem NFA.eval_append_singleton {α : Type u} {σ : Type v} (M : NFA α σ) (x : List α) (a : α) :
          M.eval (x ++ [a]) = M.stepSet (M.eval x) a
          def NFA.accepts {α : Type u} {σ : Type v} (M : NFA α σ) :

          M.accepts is the language of x such that there is an accept state in M.eval x.

          Equations
          Instances For
            theorem NFA.mem_accepts {α : Type u} {σ : Type v} {M : NFA α σ} {x : List α} :
            x M.accepts SM.accept, S M.evalFrom M.start x
            inductive NFA.Path {α : Type u} {σ : Type v} (M : NFA α σ) :
            σσList αType (max u v)

            M.Path represents a concrete path through the NFA from a start state to an end state for a particular word.

            Note that due to the non-deterministic nature of the automata, there can be more than one Path for a given word.

            Also note that this is Type and not a Prop, so that we can speak about the properties of a particular Path, such as the set of states visited along the way (defined as Path.supp).

            Instances For
              def NFA.Path.supp {α : Type u} {σ : Type v} {M : NFA α σ} [DecidableEq σ] {s t : σ} {x : List α} :
              M.Path s t xFinset σ

              Set of states visited by a path.

              Equations
              Instances For
                theorem NFA.mem_evalFrom_iff_nonempty_path {α : Type u} {σ : Type v} {M : NFA α σ} {s t : σ} {x : List α} :
                t M.evalFrom {s} x Nonempty (M.Path s t x)
                theorem NFA.accepts_iff_exists_path {α : Type u} {σ : Type v} {M : NFA α σ} {x : List α} :
                x M.accepts sM.start, tM.accept, Nonempty (M.Path s t x)
                def NFA.toDFA {α : Type u} {σ : Type v} (M : NFA α σ) :
                DFA α (Set σ)

                M.toDFA is a DFA constructed from an NFA M using the subset construction. The states is the type of Sets of M.state and the step function is M.stepSet.

                Equations
                Instances For
                  @[simp]
                  theorem NFA.toDFA_correct {α : Type u} {σ : Type v} {M : NFA α σ} :
                  theorem NFA.pumping_lemma {α : Type u} {σ : Type v} {M : NFA α σ} [Fintype σ] {x : List α} (hx : x M.accepts) (hlen : Fintype.card (Set σ) x.length) :
                  ∃ (a : List α) (b : List α) (c : List α), x = a ++ b ++ c a.length + b.length Fintype.card (Set σ) b [] {a} * KStar.kstar {b} * {c} M.accepts
                  def DFA.toNFA {α : Type u} {σ : Type v} (M : DFA α σ) :
                  NFA α σ

                  M.toNFA is an NFA constructed from a DFA M by using the same start and accept states and a transition function which sends s with input a to the singleton M.step s a.

                  Equations
                  Instances For
                    @[simp]
                    theorem DFA.toNFA_step {α : Type u} {σ : Type v} (M : DFA α σ) (s : σ) (a : α) :
                    M.toNFA.step s a = {M.step s a}
                    @[simp]
                    theorem DFA.toNFA_accept {α : Type u} {σ : Type v} (M : DFA α σ) :
                    @[simp]
                    theorem DFA.toNFA_start {α : Type u} {σ : Type v} (M : DFA α σ) :
                    @[simp]
                    theorem DFA.toNFA_evalFrom_match {α : Type u} {σ : Type v} (M : DFA α σ) (start : σ) (s : List α) :
                    M.toNFA.evalFrom {start} s = {M.evalFrom start s}
                    @[simp]
                    theorem DFA.toNFA_correct {α : Type u} {σ : Type v} (M : DFA α σ) :
                    def NFA.reverse {α : Type u} {σ : Type v} (M : NFA α σ) :
                    NFA α σ

                    M.reverse constructs an NFA with the same states as M, but all the transitions reversed. The resulting automaton accepts a word x if and only if M accepts List.reverse x.

                    Equations
                    Instances For
                      @[simp]
                      theorem NFA.reverse_accept {α : Type u} {σ : Type v} (M : NFA α σ) :
                      @[simp]
                      theorem NFA.reverse_start {α : Type u} {σ : Type v} (M : NFA α σ) :
                      @[simp]
                      theorem NFA.reverse_step {α : Type u} {σ : Type v} (M : NFA α σ) (s : σ) (a : α) :
                      M.reverse.step s a = {s' : σ | s M.step s' a}
                      @[simp]
                      theorem NFA.reverse_reverse {α : Type u} {σ : Type v} (M : NFA α σ) :
                      theorem NFA.disjoint_stepSet_reverse {α : Type u} {σ : Type v} {M : NFA α σ} {a : α} {S S' : Set σ} :
                      Disjoint S (M.reverse.stepSet S' a) Disjoint S' (M.stepSet S a)
                      theorem NFA.disjoint_evalFrom_reverse {α : Type u} {σ : Type v} {M : NFA α σ} {x : List α} {S S' : Set σ} (h : Disjoint S (M.reverse.evalFrom S' x)) :
                      theorem NFA.disjoint_evalFrom_reverse_iff {α : Type u} {σ : Type v} {M : NFA α σ} {x : List α} {S S' : Set σ} :
                      @[simp]
                      theorem NFA.mem_accepts_reverse {α : Type u} {σ : Type v} {M : NFA α σ} {x : List α} :
                      @[simp]

                      Regular languages are closed under reversal.