doc-src/TutorialI/CTL/PDL.thy
author nipkow
Fri, 06 Oct 2000 12:31:53 +0200
changeset 10159 a72ddfdbfca0
parent 10149 7cfdf6a330a0
child 10171 59d6633835fa
permissions -rw-r--r--
*** empty log message ***
     1 (*<*)theory PDL = Base:(*>*)
     2 
     3 subsection{*Propositional dynmic logic---PDL*}
     4 
     5 text{*
     6 The formulae of PDL are built up from atomic propositions via the customary
     7 propositional connectives of negation and conjunction and the two temporal
     8 connectives @{text AX} and @{text EF}. Since formulae are essentially
     9 (syntax) trees, they are naturally modelled as a datatype:
    10 *}
    11 
    12 datatype formula = Atom atom
    13                   | Neg formula
    14                   | And formula formula
    15                   | AX formula
    16                   | EF formula
    17 
    18 text{*\noindent
    19 This is almost the same as in the boolean expression case study in
    20 \S\ref{sec:boolex}, except that what used to be called @{text Var} is now
    21 called @{term Atom}.
    22 
    23 The meaning of these formulae is given by saying which formula is true in
    24 which state:
    25 *}
    26 
    27 consts valid :: "state \<Rightarrow> formula \<Rightarrow> bool"   ("(_ \<Turnstile> _)" [80,80] 80)
    28 
    29 text{*\noindent
    30 The concrete syntax annotation allows us to write @{term"s \<Turnstile> f"} instead of
    31 @{text"valid s f"}.
    32 
    33 The definition of @{text"\<Turnstile>"} is by recursion over the syntax:
    34 *}
    35 
    36 primrec
    37 "s \<Turnstile> Atom a  = (a \<in> L s)"
    38 "s \<Turnstile> Neg f   = (\<not>(s \<Turnstile> f))"
    39 "s \<Turnstile> And f g = (s \<Turnstile> f \<and> s \<Turnstile> g)"
    40 "s \<Turnstile> AX f    = (\<forall>t. (s,t) \<in> M \<longrightarrow> t \<Turnstile> f)"
    41 "s \<Turnstile> EF f    = (\<exists>t. (s,t) \<in> M^* \<and> t \<Turnstile> f)";
    42 
    43 text{*\noindent
    44 The first three equations should be self-explanatory. The temporal formula
    45 @{term"AX f"} means that @{term f} is true in all next states whereas
    46 @{term"EF f"} means that there exists some future state in which @{term f} is
    47 true. The future is expressed via @{text"^*"}, the transitive reflexive
    48 closure. Because of reflexivity, the future includes the present.
    49 
    50 Now we come to the model checker itself. It maps a formula into the set of
    51 states where the formula is true and is defined by recursion over the syntax,
    52 too:
    53 *}
    54 
    55 consts mc :: "formula \<Rightarrow> state set";
    56 primrec
    57 "mc(Atom a)  = {s. a \<in> L s}"
    58 "mc(Neg f)   = -mc f"
    59 "mc(And f g) = mc f \<inter> mc g"
    60 "mc(AX f)    = {s. \<forall>t. (s,t) \<in> M  \<longrightarrow> t \<in> mc f}"
    61 "mc(EF f)    = lfp(\<lambda>T. mc f \<union> M^-1 ^^ T)"
    62 
    63 
    64 text{*\noindent
    65 Only the equation for @{term EF} deserves some comments. Remember that the
    66 postfix @{text"^-1"} and the infix @{text"^^"} are predefined and denote the
    67 converse of a relation and the application of a relation to a set. Thus
    68 @{term "M^-1 ^^ T"} is the set of all predecessors of @{term T} and the least
    69 fixpoint (@{term lfp}) of @{term"\<lambda>T. mc f \<union> M^-1 ^^ T"} is the least set
    70 @{term T} containing @{term"mc f"} and all predecessors of @{term T}. If you
    71 find it hard to see that @{term"mc(EF f)"} contains exactly those states from
    72 which there is a path to a state where @{term f} is true, do not worry---that
    73 will be proved in a moment.
    74 
    75 First we prove monotonicity of the function inside @{term lfp}
    76 *}
    77 
    78 lemma mono_ef: "mono(\<lambda>T. A \<union> M^-1 ^^ T)"
    79 apply(rule monoI)
    80 apply blast
    81 done
    82 
    83 text{*\noindent
    84 in order to make sure it really has a least fixpoint.
    85 
    86 Now we can relate model checking and semantics. For the @{text EF} case we need
    87 a separate lemma:
    88 *}
    89 
    90 lemma EF_lemma:
    91   "lfp(\<lambda>T. A \<union> M^-1 ^^ T) = {s. \<exists>t. (s,t) \<in> M^* \<and> t \<in> A}"
    92 
    93 txt{*\noindent
    94 The equality is proved in the canonical fashion by proving that each set
    95 contains the other; the containment is shown pointwise:
    96 *}
    97 
    98 apply(rule equalityI);
    99  apply(rule subsetI);
   100  apply(simp)
   101 (*pr(latex xsymbols symbols);*)
   102 txt{*\noindent
   103 Simplification leaves us with the following first subgoal
   104 \begin{isabelle}
   105 \ \isadigit{1}{\isachardot}\ {\isasymAnd}s{\isachardot}\ s\ {\isasymin}\ lfp\ {\isacharparenleft}{\isasymlambda}T{\isachardot}\ A\ {\isasymunion}\ M{\isacharcircum}{\isacharminus}\isadigit{1}\ {\isacharcircum}{\isacharcircum}\ T{\isacharparenright}\ {\isasymLongrightarrow}\ {\isasymexists}t{\isachardot}\ {\isacharparenleft}s{\isacharcomma}\ t{\isacharparenright}\ {\isasymin}\ M{\isacharcircum}{\isacharasterisk}\ {\isasymand}\ t\ {\isasymin}\ A
   106 \end{isabelle}
   107 which is proved by @{term lfp}-induction:
   108 *}
   109 
   110  apply(erule Lfp.induct)
   111   apply(rule mono_ef)
   112  apply(simp)
   113 (*pr(latex xsymbols symbols);*)
   114 txt{*\noindent
   115 Having disposed of the monotonicity subgoal,
   116 simplification leaves us with the following main goal
   117 \begin{isabelle}
   118 \ \isadigit{1}{\isachardot}\ {\isasymAnd}s{\isachardot}\ s\ {\isasymin}\ A\ {\isasymor}\isanewline
   119 \ \ \ \ \ \ \ \ \ 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
   120 \ \ \ \ \ \ \ \ \ {\isasymLongrightarrow}\ {\isasymexists}t{\isachardot}\ {\isacharparenleft}s{\isacharcomma}\ t{\isacharparenright}\ {\isasymin}\ M{\isacharcircum}{\isacharasterisk}\ {\isasymand}\ t\ {\isasymin}\ A
   121 \end{isabelle}
   122 which is proved by @{text blast} with the help of a few lemmas about
   123 @{text"^*"}:
   124 *}
   125 
   126  apply(blast intro: r_into_rtrancl rtrancl_trans);
   127 
   128 txt{*
   129 We now return to the second set containment subgoal, which is again proved
   130 pointwise:
   131 *}
   132 
   133 apply(rule subsetI)
   134 apply(simp, clarify)
   135 (*pr(latex xsymbols symbols);*)
   136 txt{*\noindent
   137 After simplification and clarification we are left with
   138 \begin{isabelle}
   139 \ \isadigit{1}{\isachardot}\ {\isasymAnd}s\ t{\isachardot}\ {\isasymlbrakk}{\isacharparenleft}s{\isacharcomma}\ t{\isacharparenright}\ {\isasymin}\ M{\isacharcircum}{\isacharasterisk}{\isacharsemicolon}\ t\ {\isasymin}\ A{\isasymrbrakk}\ {\isasymLongrightarrow}\ s\ {\isasymin}\ lfp\ {\isacharparenleft}{\isasymlambda}T{\isachardot}\ A\ {\isasymunion}\ M{\isacharcircum}{\isacharminus}\isadigit{1}\ {\isacharcircum}{\isacharcircum}\ T{\isacharparenright}
   140 \end{isabelle}
   141 This goal is proved by induction on @{term"(s,t)\<in>M^*"}. But since the model
   142 checker works backwards (from @{term t} to @{term s}), we cannot use the
   143 induction theorem @{thm[source]rtrancl_induct} because it works in the
   144 forward direction. Fortunately the converse induction theorem
   145 @{thm[source]converse_rtrancl_induct} already exists:
   146 @{thm[display,margin=60]converse_rtrancl_induct[no_vars]}
   147 It says that if @{prop"(a,b):r^*"} and we know @{prop"P b"} then we can infer
   148 @{prop"P a"} provided each step backwards from a predecessor @{term z} of
   149 @{term b} preserves @{term P}.
   150 *}
   151 
   152 apply(erule converse_rtrancl_induct)
   153 (*pr(latex xsymbols symbols);*)
   154 txt{*\noindent
   155 The base case
   156 \begin{isabelle}
   157 \ \isadigit{1}{\isachardot}\ {\isasymAnd}t{\isachardot}\ t\ {\isasymin}\ A\ {\isasymLongrightarrow}\ t\ {\isasymin}\ lfp\ {\isacharparenleft}{\isasymlambda}T{\isachardot}\ A\ {\isasymunion}\ M{\isacharcircum}{\isacharminus}\isadigit{1}\ {\isacharcircum}{\isacharcircum}\ T{\isacharparenright}
   158 \end{isabelle}
   159 is solved by unrolling @{term lfp} once
   160 *}
   161 
   162  apply(rule ssubst[OF lfp_Tarski[OF mono_ef]])
   163 (*pr(latex xsymbols symbols);*)
   164 txt{*
   165 \begin{isabelle}
   166 \ \isadigit{1}{\isachardot}\ {\isasymAnd}t{\isachardot}\ t\ {\isasymin}\ A\ {\isasymLongrightarrow}\ t\ {\isasymin}\ A\ {\isasymunion}\ M{\isacharcircum}{\isacharminus}\isadigit{1}\ {\isacharcircum}{\isacharcircum}\ lfp\ {\isacharparenleft}{\isasymlambda}T{\isachardot}\ A\ {\isasymunion}\ M{\isacharcircum}{\isacharminus}\isadigit{1}\ {\isacharcircum}{\isacharcircum}\ T{\isacharparenright}
   167 \end{isabelle}
   168 and disposing of the resulting trivial subgoal automatically:
   169 *}
   170 
   171  apply(blast)
   172 
   173 txt{*\noindent
   174 The proof of the induction step is identical to the one for the base case:
   175 *}
   176 
   177 apply(rule ssubst[OF lfp_Tarski[OF mono_ef]])
   178 apply(blast)
   179 done
   180 
   181 text{*
   182 The main theorem is proved in the familiar manner: induction followed by
   183 @{text auto} augmented with the lemma as a simplification rule.
   184 *}
   185 
   186 theorem "mc f = {s. s \<Turnstile> f}";
   187 apply(induct_tac f);
   188 apply(auto simp add:EF_lemma);
   189 done;
   190 (*<*)end(*>*)