Documentation

Cslib.Computability.Automata.DA

Deterministic Automata #

A Deterministic Automaton (DA) is an automaton defined by a transition function equipped with an initial state.

Automata with different accepting conditions are then defined as extensions of DA. These include, for example, a generalised version of DFA as found in the literature (without finiteness assumptions), deterministic Buchi automata, and deterministic Muller automata.

References #

structure Cslib.Automata.DA (State : Type u_1) (Symbol : Type u_2) extends Cslib.FLTS State Symbol :
Type (max u_1 u_2)

A deterministic automaton extends a FLTS with a unique initial state.

  • tr : StateSymbolState
  • start : State

    The initial state of the automaton.

Instances For
    def Cslib.Automata.DA.run' {State : Type u_1} {Symbol : Type u_2} (da : DA State Symbol) (xs : ωSequence Symbol) :
    State

    Helper function for defining run below.

    Equations
    Instances For
      def Cslib.Automata.DA.run {State : Type u_1} {Symbol : Type u_2} (da : DA State Symbol) (xs : ωSequence Symbol) :

      Infinite run.

      Equations
      Instances For
        @[simp]
        theorem Cslib.Automata.DA.run_zero {State : Type u_1} {Symbol : Type u_2} {da : DA State Symbol} {xs : ωSequence Symbol} :
        (da.run xs) 0 = da.start
        @[simp]
        theorem Cslib.Automata.DA.run_succ {State : Type u_1} {Symbol : Type u_2} {da : DA State Symbol} {xs : ωSequence Symbol} {n : } :
        (da.run xs) (n + 1) = da.tr ((da.run xs) n) (xs n)
        @[simp]
        theorem Cslib.Automata.DA.mtr_extract_eq_run {State : Type u_1} {Symbol : Type u_2} {da : DA State Symbol} {xs : ωSequence Symbol} {n : } :
        da.mtr da.start (xs.extract 0 n) = (da.run xs) n
        structure Cslib.Automata.DA.FinAcc (State : Type u_3) (Symbol : Type u_4) extends Cslib.Automata.DA State Symbol :
        Type (max u_3 u_4)

        A deterministic automaton that accepts finite strings (lists of symbols).

        • tr : StateSymbolState
        • start : State
        • accept : Set State

          The set of accepting states.

        Instances For
          instance Cslib.Automata.DA.FinAcc.instAcceptor {State : Type u_1} {Symbol : Type u_2} :
          Acceptor (FinAcc State Symbol) Symbol

          A DA.FinAcc accepts a string if its multistep transition function maps the start state and the string to an accept state.

          This is the standard string recognition performed by DFAs in the literature.

          Equations
          structure Cslib.Automata.DA.Buchi (State : Type u_3) (Symbol : Type u_4) extends Cslib.Automata.DA State Symbol :
          Type (max u_3 u_4)

          Deterministic Buchi automaton.

          • tr : StateSymbolState
          • start : State
          • accept : Set State

            The set of accepting states.

          Instances For
            instance Cslib.Automata.DA.Buchi.instωAcceptor {State : Type u_1} {Symbol : Type u_2} :
            ωAcceptor (Buchi State Symbol) Symbol

            An infinite run is accepting iff accepting states occur infinitely many times.

            Equations
            structure Cslib.Automata.DA.Muller (State : Type u_3) (Symbol : Type u_4) extends Cslib.Automata.DA State Symbol :
            Type (max u_3 u_4)

            Deterministic Muller automaton.

            • tr : StateSymbolState
            • start : State
            • accept : Set (Set State)

              The set of sets of accepting states.

            Instances For
              instance Cslib.Automata.DA.Muller.instωAcceptor {State : Type u_1} {Symbol : Type u_2} :
              ωAcceptor (Muller State Symbol) Symbol

              An infinite run is accepting iff the set of states that occur infinitely many times is one of the sets in accept.

              Equations