diff -r a5f53d9d2b60 -r f8537d69f514 doc-src/TutorialI/CTL/PDL.thy --- a/doc-src/TutorialI/CTL/PDL.thy Thu May 29 13:27:13 2008 +0200 +++ b/doc-src/TutorialI/CTL/PDL.thy Thu May 29 22:45:33 2008 +0200 @@ -21,23 +21,17 @@ text{*\noindent This resembles the boolean expression case study in \S\ref{sec:boolex}. -A validity relation between -states and formulae specifies the semantics: +A validity relation between states and formulae specifies the semantics. +The syntax annotation allows us to write @{text"s \ f"} instead of +\hbox{@{text"valid s f"}}. The definition is by recursion over the syntax: *} -consts valid :: "state \ formula \ bool" ("(_ \ _)" [80,80] 80) - -text{*\noindent -The syntax annotation allows us to write @{term"s \ f"} instead of -\hbox{@{text"valid s f"}}. -The definition of @{text"\"} is by recursion over the syntax: -*} - -primrec -"s \ Atom a = (a \ L s)" -"s \ Neg f = (\(s \ f))" -"s \ And f g = (s \ f \ s \ g)" -"s \ AX f = (\t. (s,t) \ M \ t \ f)" +primrec valid :: "state \ formula \ bool" ("(_ \ _)" [80,80] 80) +where +"s \ Atom a = (a \ L s)" | +"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\<^sup>* \ t \ f)" text{*\noindent @@ -47,16 +41,15 @@ true. The future is expressed via @{text"\<^sup>*"}, the reflexive transitive 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 -states where the formula is true. It too is defined by recursion over the syntax: -*} +Now we come to the model checker itself. It maps a formula into the +set of states where the formula is true. It too is defined by +recursion over the syntax: *} -consts mc :: "formula \ state set" -primrec -"mc(Atom a) = {s. a \ L s}" -"mc(Neg f) = -mc f" -"mc(And f g) = mc f \ mc g" -"mc(AX f) = {s. \t. (s,t) \ M \ t \ mc f}" +primrec mc :: "formula \ state set" where +"mc(Atom a) = {s. a \ L s}" | +"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\ `` T))" text{*\noindent