Theory Ta
section "Tree Automata"
theory Ta
imports Main Automatic_Refinement.Misc Tree
begin
text_raw ‹\label{sec:ta}›
text ‹
This theory defines tree automata, tree regular languages and
specifies basic algorithms.
Nondeterministic and deterministic (bottom-up) tree automata are defined.
For non-deterministic tree automata, basic algorithms for membership, union,
intersection, forward and backward reduction, and emptiness check are
specified.
Moreover, a (brute-force) determinization algorithm is specified.
For deterministic tree automata, we specify algorithms for complement
and completion.
Finally, the class of regular languages over a given ranked alphabet is defined
and its standard closure properties are proved.
The specification of the algorithms in this theory is very high-level, and the
specifications are not executable. A bit more specific algorithms are defined
in Section~\ref{sec:absalgo}, and a refinement to executable definitions is
done in Section~\ref{sec:taimpl}.
›
subsection "Basic Definitions"
subsubsection "Tree Automata"
text ‹
A tree automata consists of a (finite) set of initial states
and a (finite) set of rules.
A rule has the form ‹q → l q1…qn›,
with the meaning that one can derive
‹l(q1…qn)› from the state ‹q›.
›