diff -r 2d4c058749a7 -r c00ac928fc6f doc-src/TutorialI/CTL/PDL.thy --- a/doc-src/TutorialI/CTL/PDL.thy Sat Jan 06 11:27:09 2001 +0100 +++ b/doc-src/TutorialI/CTL/PDL.thy Sat Jan 06 12:39:05 2001 +0100 @@ -38,13 +38,13 @@ "s \ Neg f = (\(s \ f))" "s \ And f g = (s \ f \ s \ g)" "s \ AX f = (\t. (s,t) \ M \ t \ f)" -"s \ EF f = (\t. (s,t) \ M^* \ t \ f)"; +"s \ EF f = (\t. (s,t) \ M\<^sup>* \ t \ f)"; text{*\noindent The first three equations should be self-explanatory. The temporal formula @{term"AX f"} means that @{term f} is true in all next states whereas @{term"EF f"} means that there exists some future state in which @{term f} is -true. The future is expressed via @{text"^*"}, the transitive reflexive +true. The future is expressed via @{text"\<^sup>*"}, the transitive reflexive closure. Because of reflexivity, the future includes the present. Now we come to the model checker itself. It maps a formula into the set of @@ -58,14 +58,14 @@ "mc(Neg f) = -mc f" "mc(And f g) = mc f \ mc g" "mc(AX f) = {s. \t. (s,t) \ M \ t \ mc f}" -"mc(EF f) = lfp(\T. mc f \ M^-1 ``` T)" +"mc(EF f) = lfp(\T. mc f \ M\ ``` T)" text{*\noindent Only the equation for @{term EF} deserves some comments. Remember that the -postfix @{text"^-1"} and the infix @{text"```"} are predefined and denote the +postfix @{text"\"} and the infix @{text"```"} are predefined and denote the converse of a relation and the application of a relation to a set. Thus -@{term "M^-1 ``` T"} is the set of all predecessors of @{term T} and the least -fixed point (@{term lfp}) of @{term"\T. mc f \ M^-1 ``` T"} is the least set +@{term "M\ ``` T"} is the set of all predecessors of @{term T} and the least +fixed point (@{term lfp}) of @{term"\T. mc f \ M\ ``` T"} is the least set @{term T} containing @{term"mc f"} and all predecessors of @{term T}. If you find it hard to see that @{term"mc(EF f)"} contains exactly those states from which there is a path to a state where @{term f} is true, do not worry---that @@ -74,7 +74,7 @@ First we prove monotonicity of the function inside @{term lfp} *} -lemma mono_ef: "mono(\T. A \ M^-1 ``` T)" +lemma mono_ef: "mono(\T. A \ M\ ``` T)" apply(rule monoI) apply blast done @@ -87,7 +87,7 @@ *} lemma EF_lemma: - "lfp(\T. A \ M^-1 ``` T) = {s. \t. (s,t) \ M^* \ t \ A}" + "lfp(\T. A \ M\ ``` T) = {s. \t. (s,t) \ M\<^sup>* \ t \ A}" txt{*\noindent The equality is proved in the canonical fashion by proving that each set @@ -112,11 +112,11 @@ Having disposed of the monotonicity subgoal, simplification leaves us with the following main goal \begin{isabelle} -\ \isadigit{1}{\isachardot}\ {\isasymAnd}s{\isachardot}\ s\ {\isasymin}\ A\ {\isasymor}\isanewline -\ \ \ \ \ \ \ \ \ s\ {\isasymin}\ M{\isacharcircum}{\isacharminus}\isadigit{1}\ {\isacharcircum}{\isacharcircum}\ {\isacharparenleft}lfp\ {\isacharparenleft}{\dots}{\isacharparenright}\ {\isasyminter}\ {\isacharbraceleft}x{\isachardot}\ {\isasymexists}t{\isachardot}\ {\isacharparenleft}x{\isacharcomma}\ t{\isacharparenright}\ {\isasymin}\ M{\isacharcircum}{\isacharasterisk}\ {\isasymand}\ t\ {\isasymin}\ A{\isacharbraceright}{\isacharparenright}\isanewline -\ \ \ \ \ \ \ \ \ {\isasymLongrightarrow}\ {\isasymexists}t{\isachardot}\ {\isacharparenleft}s{\isacharcomma}\ t{\isacharparenright}\ {\isasymin}\ M{\isacharcircum}{\isacharasterisk}\ {\isasymand}\ t\ {\isasymin}\ A +\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}x{\isachardot}\ x\ {\isasymin}\ A\ {\isasymor}\isanewline +\ \ \ \ \ \ \ \ \ x\ {\isasymin}\ M{\isasyminverse}\ {\isacharbackquote}{\isacharbackquote}{\isacharbackquote}\ {\isacharparenleft}lfp\ {\isacharparenleft}\dots{\isacharparenright}\ {\isasyminter}\ {\isacharbraceleft}x{\isachardot}\ {\isasymexists}t{\isachardot}\ {\isacharparenleft}x{\isacharcomma}\ t{\isacharparenright}\ {\isasymin}\ M\isactrlsup {\isacharasterisk}\ {\isasymand}\ t\ {\isasymin}\ A{\isacharbraceright}{\isacharparenright}\isanewline +\ \ \ \ \ \ \ \ {\isasymLongrightarrow}\ {\isasymexists}t{\isachardot}\ {\isacharparenleft}x{\isacharcomma}\ t{\isacharparenright}\ {\isasymin}\ M\isactrlsup {\isacharasterisk}\ {\isasymand}\ t\ {\isasymin}\ A \end{isabelle} -which is proved by @{text blast} with the help of transitivity of @{text"^*"}: +which is proved by @{text blast} with the help of transitivity of @{text"\<^sup>*"}: *} apply(blast intro: rtrancl_trans); @@ -132,13 +132,13 @@ txt{*\noindent After simplification and clarification we are left with @{subgoals[display,indent=0,goals_limit=1]} -This goal is proved by induction on @{term"(s,t)\M^*"}. But since the model +This goal is proved by induction on @{term"(s,t)\M\<^sup>*"}. But since the model checker works backwards (from @{term t} to @{term s}), we cannot use the induction theorem @{thm[source]rtrancl_induct} because it works in the forward direction. Fortunately the converse induction theorem @{thm[source]converse_rtrancl_induct} already exists: @{thm[display,margin=60]converse_rtrancl_induct[no_vars]} -It says that if @{prop"(a,b):r^*"} and we know @{prop"P b"} then we can infer +It says that if @{prop"(a,b):r\<^sup>*"} and we know @{prop"P b"} then we can infer @{prop"P a"} provided each step backwards from a predecessor @{term z} of @{term b} preserves @{term P}. *}