diff -r dfff821d2949 -r 59d6633835fa doc-src/TutorialI/CTL/PDL.thy --- a/doc-src/TutorialI/CTL/PDL.thy Mon Oct 09 09:33:45 2000 +0200 +++ b/doc-src/TutorialI/CTL/PDL.thy Mon Oct 09 10:18:21 2000 +0200 @@ -187,4 +187,34 @@ apply(induct_tac f); apply(auto simp add:EF_lemma); done; -(*<*)end(*>*) + +text{* +\begin{exercise} +@{term AX} has a dual operator @{term EN}\footnote{We cannot use the customary @{text EX} +as that is the ASCII equivalent of @{text"\"}} +(``there exists a next state such that'') with the intended semantics +@{prop[display]"(s \ EN f) = (EX t. (s,t) : M & t \ f)"} +Fortunately, @{term"EN f"} can already be expressed as a PDL formula. How? + +Show that the semantics for @{term EF} satisfies the following recursion equation: +@{prop[display]"(s \ EF f) = (s \ f | s \ EN(EF f))"} +\end{exercise} +*} +(*<*) +theorem main: "mc f = {s. s \ f}"; +apply(induct_tac f); +apply(auto simp add:EF_lemma); +done; + +lemma aux: "s \ f = (s : mc f)"; +apply(simp add:main); +done; + +lemma "(s \ EF f) = (s \ f | s \ Neg(AX(Neg(EF f))))"; +apply(simp only:aux); +apply(simp); +apply(rule ssubst[OF lfp_Tarski[OF mono_ef]], fast); +done + +end +(*>*)