Theory DGBTA

section ‹Deterministic Generalized Büchi Transition Automata›

theory DGBTA
imports "../Deterministic"
begin

  datatype ('label, 'state) dgbta = dgbta
    (alphabet: "'label set")
    (initial: "'state")
    (transition: "'label  'state  'state")
    (accepting: "('state × 'label × 'state) pred gen")

  global_interpretation dgbta: automaton dgbta alphabet initial transition accepting
    defines path = dgbta.path and run = dgbta.run and reachable = dgbta.reachable and nodes = dgbta.nodes
    by unfold_locales auto
  global_interpretation dgbta: automaton_run dgbta alphabet initial transition accepting
    "λ P w r p. gen infs P (p ## r ||| w ||| r)"
    defines language = dgbta.language
    by standard

  abbreviation target where "target  dgbta.target"
  abbreviation states where "states  dgbta.states"
  abbreviation trace where "trace  dgbta.trace"
  abbreviation successors where "successors  dgbta.successors TYPE('label)"

end