Theory LTL.LTL
section ‹Linear Temporal Logic›
theory LTL
imports
Main "HOL-Library.Omega_Words_Fun"
begin
text ‹This theory provides a formalisation of linear temporal logic. It provides three variants:
\begin{enumerate}
\item LTL with syntactic sugar. This variant is the semantic reference and the included parser
generates ASTs of this datatype.
\item LTL in negation normal form without syntactic sugar. This variant is used by the included
rewriting engine and is used for the translation to automata (implemented in other entries).
\item LTL in restricted negation normal form without the rather uncommon operators ``weak until''
and ``strong release''. It is used by the formalization of Gerth's algorithm.
\item PLTL. A variant with a reduced set of operators.
\end{enumerate}
This theory subsumes (and partly reuses) the existing formalisation found in LTL\_to\_GBA and
Stuttering\_Equivalence and unifies them.›
subsection ‹LTL with Syntactic Sugar›
text ‹In this section, we provide a formulation of LTL with explicit syntactic sugar deeply embedded.
This formalization serves as a reference semantics.›
subsubsection ‹Syntax›