Theory TLS
theory TLS
imports
"Safety_Logic"
begin
section‹ A Temporal Logic of Safety (TLS) \label{sec:tls} ›
text‹
We model systems with finite and infinite sequences of states, closed
under stuttering following \<^citet>‹"Lamport:1994"›. This theory relates
the safety logic of \S\ref{sec:safety_logic} to the powerset
(quotiented by stuttering) representing properties of these sequences
(see \S\ref{sec:tls-safety}). Most of this story is standard but the
addition of finite sequences does have some impact.
References:
▪ historical motivations for future-time linear temporal logic (LTL): \<^citet>‹"MannaPnueli:1991" and "OwickiLamport:1982"›.
▪ a discussion on the merits of proving liveness: 🌐‹https://cs.nyu.edu/acsys/beyond-safety/liveness.htm›
Observations:
▪ Lamport (and Abadi et al) treat infinite stuttering as termination
▪ \<^citet>‹‹p189› in "Lamport:1999"›: ``we can represent a terminating execution of any system by an
infinite behavior that ends with a sequence of nothing but stuttering steps. We have no need of
finite behaviors (finite sequences of states), so we consider only infinite ones.''
▪ this conflates divergence with termination
▪ we separate those concepts here so we can support sequential composition
▪ the traditional account of liveness properties breaks down (see \S\ref{sec:safety_closure})
›
subsection‹ Stuttering\label{sec:tls-stuttering} ›
text‹
An infinitary version of \<^const>‹trace.natural'›.
Observations:
▪ we need to normalize the agent labels for sequences that infinitely stutter
Source materials:
▪ 🗏‹$ISABELLE_HOME/src/HOL/Corec_Examples/LFilter.thy›.
▪ ▩‹$AFP/Coinductive/Coinductive_List.thy›
▪ ▩‹$AFP/Coinductive/TLList.thy›
▪ ▩‹$AFP/TLA/Sequence.thy›.
›
definition trailing :: "'c ⇒ ('a, 'b) tllist ⇒ ('c, 'b) tllist" where
"trailing s xs = (if tfinite xs then TNil (terminal xs) else trepeat s)"