doc-src/TutorialI/CTL/PDL.thy
changeset 10801 c00ac928fc6f
parent 10800 2d4c058749a7
child 10839 1f93f5a27de6
     1.1 --- a/doc-src/TutorialI/CTL/PDL.thy	Sat Jan 06 11:27:09 2001 +0100
     1.2 +++ b/doc-src/TutorialI/CTL/PDL.thy	Sat Jan 06 12:39:05 2001 +0100
     1.3 @@ -38,13 +38,13 @@
     1.4  "s \<Turnstile> Neg f   = (\<not>(s \<Turnstile> f))"
     1.5  "s \<Turnstile> And f g = (s \<Turnstile> f \<and> s \<Turnstile> g)"
     1.6  "s \<Turnstile> AX f    = (\<forall>t. (s,t) \<in> M \<longrightarrow> t \<Turnstile> f)"
     1.7 -"s \<Turnstile> EF f    = (\<exists>t. (s,t) \<in> M^* \<and> t \<Turnstile> f)";
     1.8 +"s \<Turnstile> EF f    = (\<exists>t. (s,t) \<in> M\<^sup>* \<and> t \<Turnstile> f)";
     1.9  
    1.10  text{*\noindent
    1.11  The first three equations should be self-explanatory. The temporal formula
    1.12  @{term"AX f"} means that @{term f} is true in all next states whereas
    1.13  @{term"EF f"} means that there exists some future state in which @{term f} is
    1.14 -true. The future is expressed via @{text"^*"}, the transitive reflexive
    1.15 +true. The future is expressed via @{text"\<^sup>*"}, the transitive reflexive
    1.16  closure. Because of reflexivity, the future includes the present.
    1.17  
    1.18  Now we come to the model checker itself. It maps a formula into the set of
    1.19 @@ -58,14 +58,14 @@
    1.20  "mc(Neg f)   = -mc f"
    1.21  "mc(And f g) = mc f \<inter> mc g"
    1.22  "mc(AX f)    = {s. \<forall>t. (s,t) \<in> M  \<longrightarrow> t \<in> mc f}"
    1.23 -"mc(EF f)    = lfp(\<lambda>T. mc f \<union> M^-1 ``` T)"
    1.24 +"mc(EF f)    = lfp(\<lambda>T. mc f \<union> M\<inverse> ``` T)"
    1.25  
    1.26  text{*\noindent
    1.27  Only the equation for @{term EF} deserves some comments. Remember that the
    1.28 -postfix @{text"^-1"} and the infix @{text"```"} are predefined and denote the
    1.29 +postfix @{text"\<inverse>"} and the infix @{text"```"} are predefined and denote the
    1.30  converse of a relation and the application of a relation to a set. Thus
    1.31 -@{term "M^-1 ``` T"} is the set of all predecessors of @{term T} and the least
    1.32 -fixed point (@{term lfp}) of @{term"\<lambda>T. mc f \<union> M^-1 ``` T"} is the least set
    1.33 +@{term "M\<inverse> ``` T"} is the set of all predecessors of @{term T} and the least
    1.34 +fixed point (@{term lfp}) of @{term"\<lambda>T. mc f \<union> M\<inverse> ``` T"} is the least set
    1.35  @{term T} containing @{term"mc f"} and all predecessors of @{term T}. If you
    1.36  find it hard to see that @{term"mc(EF f)"} contains exactly those states from
    1.37  which there is a path to a state where @{term f} is true, do not worry---that
    1.38 @@ -74,7 +74,7 @@
    1.39  First we prove monotonicity of the function inside @{term lfp}
    1.40  *}
    1.41  
    1.42 -lemma mono_ef: "mono(\<lambda>T. A \<union> M^-1 ``` T)"
    1.43 +lemma mono_ef: "mono(\<lambda>T. A \<union> M\<inverse> ``` T)"
    1.44  apply(rule monoI)
    1.45  apply blast
    1.46  done
    1.47 @@ -87,7 +87,7 @@
    1.48  *}
    1.49  
    1.50  lemma EF_lemma:
    1.51 -  "lfp(\<lambda>T. A \<union> M^-1 ``` T) = {s. \<exists>t. (s,t) \<in> M^* \<and> t \<in> A}"
    1.52 +  "lfp(\<lambda>T. A \<union> M\<inverse> ``` T) = {s. \<exists>t. (s,t) \<in> M\<^sup>* \<and> t \<in> A}"
    1.53  
    1.54  txt{*\noindent
    1.55  The equality is proved in the canonical fashion by proving that each set
    1.56 @@ -112,11 +112,11 @@
    1.57  Having disposed of the monotonicity subgoal,
    1.58  simplification leaves us with the following main goal
    1.59  \begin{isabelle}
    1.60 -\ \isadigit{1}{\isachardot}\ {\isasymAnd}s{\isachardot}\ s\ {\isasymin}\ A\ {\isasymor}\isanewline
    1.61 -\ \ \ \ \ \ \ \ \ 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
    1.62 -\ \ \ \ \ \ \ \ \ {\isasymLongrightarrow}\ {\isasymexists}t{\isachardot}\ {\isacharparenleft}s{\isacharcomma}\ t{\isacharparenright}\ {\isasymin}\ M{\isacharcircum}{\isacharasterisk}\ {\isasymand}\ t\ {\isasymin}\ A
    1.63 +\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}x{\isachardot}\ x\ {\isasymin}\ A\ {\isasymor}\isanewline
    1.64 +\ \ \ \ \ \ \ \ \ 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
    1.65 +\ \ \ \ \ \ \ \ {\isasymLongrightarrow}\ {\isasymexists}t{\isachardot}\ {\isacharparenleft}x{\isacharcomma}\ t{\isacharparenright}\ {\isasymin}\ M\isactrlsup {\isacharasterisk}\ {\isasymand}\ t\ {\isasymin}\ A
    1.66  \end{isabelle}
    1.67 -which is proved by @{text blast} with the help of transitivity of @{text"^*"}:
    1.68 +which is proved by @{text blast} with the help of transitivity of @{text"\<^sup>*"}:
    1.69  *}
    1.70  
    1.71   apply(blast intro: rtrancl_trans);
    1.72 @@ -132,13 +132,13 @@
    1.73  txt{*\noindent
    1.74  After simplification and clarification we are left with
    1.75  @{subgoals[display,indent=0,goals_limit=1]}
    1.76 -This goal is proved by induction on @{term"(s,t)\<in>M^*"}. But since the model
    1.77 +This goal is proved by induction on @{term"(s,t)\<in>M\<^sup>*"}. But since the model
    1.78  checker works backwards (from @{term t} to @{term s}), we cannot use the
    1.79  induction theorem @{thm[source]rtrancl_induct} because it works in the
    1.80  forward direction. Fortunately the converse induction theorem
    1.81  @{thm[source]converse_rtrancl_induct} already exists:
    1.82  @{thm[display,margin=60]converse_rtrancl_induct[no_vars]}
    1.83 -It says that if @{prop"(a,b):r^*"} and we know @{prop"P b"} then we can infer
    1.84 +It says that if @{prop"(a,b):r\<^sup>*"} and we know @{prop"P b"} then we can infer
    1.85  @{prop"P a"} provided each step backwards from a predecessor @{term z} of
    1.86  @{term b} preserves @{term P}.
    1.87  *}