# HG changeset patch # User nipkow # Date 978781145 -3600 # Node ID c00ac928fc6f58ccaa9cfb375a7923f07eaf71cb # Parent 2d4c058749a74900abb8bbb08dadbb6818006f3d *** empty log message *** diff -r 2d4c058749a7 -r c00ac928fc6f doc-src/TutorialI/CTL/CTL.thy --- a/doc-src/TutorialI/CTL/CTL.thy Sat Jan 06 11:27:09 2001 +0100 +++ b/doc-src/TutorialI/CTL/CTL.thy Sat Jan 06 12:39:05 2001 +0100 @@ -39,7 +39,7 @@ "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)" (*>*) "s \ AF f = (\p \ Paths s. \i. p i \ f)"; @@ -62,7 +62,7 @@ "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)"(*>*) "mc(AF f) = lfp(af(mc f))"; text{*\noindent @@ -75,12 +75,12 @@ apply blast; done (*<*) -lemma mono_ef: "mono(\T. A \ M^-1 ``` T)"; +lemma mono_ef: "mono(\T. A \ M\ ``` T)"; apply(rule monoI); by(blast); 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}"; apply(rule equalityI); apply(rule subsetI); apply(simp); @@ -366,7 +366,7 @@ Note that @{term EU} is not definable in terms of the other operators! Model checking @{term EU} is again a least fixed point construction: -@{text[display]"mc(EU f g) = lfp(\T. mc g \ mc f \ (M^-1 ``` T))"} +@{text[display]"mc(EU f g) = lfp(\T. mc g \ mc f \ (M\ ``` T))"} \begin{exercise} Extend the datatype of formulae by the above until operator @@ -382,7 +382,7 @@ (*<*) constdefs eufix :: "state set \ state set \ state set \ state set" -"eufix A B T \ B \ A \ (M^-1 ``` T)" +"eufix A B T \ B \ A \ (M\ ``` T)" lemma "lfp(eufix A B) \ eusem A B" apply(rule lfp_lowerbound) 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}. *} diff -r 2d4c058749a7 -r c00ac928fc6f doc-src/TutorialI/CTL/document/CTL.tex --- a/doc-src/TutorialI/CTL/document/CTL.tex Sat Jan 06 11:27:09 2001 +0100 +++ b/doc-src/TutorialI/CTL/document/CTL.tex Sat Jan 06 12:39:05 2001 +0100 @@ -300,7 +300,7 @@ Model checking \isa{EU} is again a least fixed point construction: \begin{isabelle}% -\ \ \ \ \ mc{\isacharparenleft}EU\ f\ g{\isacharparenright}\ {\isacharequal}\ lfp{\isacharparenleft}{\isasymlambda}T{\isachardot}\ mc\ g\ {\isasymunion}\ mc\ f\ {\isasyminter}\ {\isacharparenleft}M{\isacharcircum}{\isacharminus}{\isadigit{1}}\ {\isacharbackquote}{\isacharbackquote}{\isacharbackquote}\ T{\isacharparenright}{\isacharparenright}% +\ \ \ \ \ mc{\isacharparenleft}EU\ f\ g{\isacharparenright}\ {\isacharequal}\ lfp{\isacharparenleft}{\isasymlambda}T{\isachardot}\ mc\ g\ {\isasymunion}\ mc\ f\ {\isasyminter}\ {\isacharparenleft}M{\isasyminverse}\ {\isacharbackquote}{\isacharbackquote}{\isacharbackquote}\ T{\isacharparenright}{\isacharparenright}% \end{isabelle} \begin{exercise} diff -r 2d4c058749a7 -r c00ac928fc6f doc-src/TutorialI/CTL/document/PDL.tex --- a/doc-src/TutorialI/CTL/document/PDL.tex Sat Jan 06 11:27:09 2001 +0100 +++ b/doc-src/TutorialI/CTL/document/PDL.tex Sat Jan 06 12:39:05 2001 +0100 @@ -39,13 +39,13 @@ {\isachardoublequote}s\ {\isasymTurnstile}\ Neg\ f\ \ \ {\isacharequal}\ {\isacharparenleft}{\isasymnot}{\isacharparenleft}s\ {\isasymTurnstile}\ f{\isacharparenright}{\isacharparenright}{\isachardoublequote}\isanewline {\isachardoublequote}s\ {\isasymTurnstile}\ And\ f\ g\ {\isacharequal}\ {\isacharparenleft}s\ {\isasymTurnstile}\ f\ {\isasymand}\ s\ {\isasymTurnstile}\ g{\isacharparenright}{\isachardoublequote}\isanewline {\isachardoublequote}s\ {\isasymTurnstile}\ AX\ f\ \ \ \ {\isacharequal}\ {\isacharparenleft}{\isasymforall}t{\isachardot}\ {\isacharparenleft}s{\isacharcomma}t{\isacharparenright}\ {\isasymin}\ M\ {\isasymlongrightarrow}\ t\ {\isasymTurnstile}\ f{\isacharparenright}{\isachardoublequote}\isanewline -{\isachardoublequote}s\ {\isasymTurnstile}\ EF\ f\ \ \ \ {\isacharequal}\ {\isacharparenleft}{\isasymexists}t{\isachardot}\ {\isacharparenleft}s{\isacharcomma}t{\isacharparenright}\ {\isasymin}\ M{\isacharcircum}{\isacharasterisk}\ {\isasymand}\ t\ {\isasymTurnstile}\ f{\isacharparenright}{\isachardoublequote}% +{\isachardoublequote}s\ {\isasymTurnstile}\ EF\ f\ \ \ \ {\isacharequal}\ {\isacharparenleft}{\isasymexists}t{\isachardot}\ {\isacharparenleft}s{\isacharcomma}t{\isacharparenright}\ {\isasymin}\ M\isactrlsup {\isacharasterisk}\ {\isasymand}\ t\ {\isasymTurnstile}\ f{\isacharparenright}{\isachardoublequote}% \begin{isamarkuptext}% \noindent The first three equations should be self-explanatory. The temporal formula \isa{AX\ f} means that \isa{f} is true in all next states whereas \isa{EF\ f} means that there exists some future state in which \isa{f} is -true. The future is expressed via \isa{{\isacharcircum}{\isacharasterisk}}, the transitive reflexive +true. The future is expressed via \isa{\isactrlsup {\isacharasterisk}}, 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,11 +58,11 @@ {\isachardoublequote}mc{\isacharparenleft}Neg\ f{\isacharparenright}\ \ \ {\isacharequal}\ {\isacharminus}mc\ f{\isachardoublequote}\isanewline {\isachardoublequote}mc{\isacharparenleft}And\ f\ g{\isacharparenright}\ {\isacharequal}\ mc\ f\ {\isasyminter}\ mc\ g{\isachardoublequote}\isanewline {\isachardoublequote}mc{\isacharparenleft}AX\ f{\isacharparenright}\ \ \ \ {\isacharequal}\ {\isacharbraceleft}s{\isachardot}\ {\isasymforall}t{\isachardot}\ {\isacharparenleft}s{\isacharcomma}t{\isacharparenright}\ {\isasymin}\ M\ \ {\isasymlongrightarrow}\ t\ {\isasymin}\ mc\ f{\isacharbraceright}{\isachardoublequote}\isanewline -{\isachardoublequote}mc{\isacharparenleft}EF\ f{\isacharparenright}\ \ \ \ {\isacharequal}\ lfp{\isacharparenleft}{\isasymlambda}T{\isachardot}\ mc\ f\ {\isasymunion}\ M{\isacharcircum}{\isacharminus}{\isadigit{1}}\ {\isacharbackquote}{\isacharbackquote}{\isacharbackquote}\ T{\isacharparenright}{\isachardoublequote}% +{\isachardoublequote}mc{\isacharparenleft}EF\ f{\isacharparenright}\ \ \ \ {\isacharequal}\ lfp{\isacharparenleft}{\isasymlambda}T{\isachardot}\ mc\ f\ {\isasymunion}\ M{\isasyminverse}\ {\isacharbackquote}{\isacharbackquote}{\isacharbackquote}\ T{\isacharparenright}{\isachardoublequote}% \begin{isamarkuptext}% \noindent Only the equation for \isa{EF} deserves some comments. Remember that the -postfix \isa{{\isacharcircum}{\isacharminus}{\isadigit{1}}} and the infix \isa{{\isacharbackquote}{\isacharbackquote}{\isacharbackquote}} are predefined and denote the +postfix \isa{{\isasyminverse}} and the infix \isa{{\isacharbackquote}{\isacharbackquote}{\isacharbackquote}} are predefined and denote the converse of a relation and the application of a relation to a set. Thus \isa{M{\isasyminverse}\ {\isacharbackquote}{\isacharbackquote}{\isacharbackquote}\ T} is the set of all predecessors of \isa{T} and the least fixed point (\isa{lfp}) of \isa{{\isasymlambda}T{\isachardot}\ mc\ f\ {\isasymunion}\ M{\isasyminverse}\ {\isacharbackquote}{\isacharbackquote}{\isacharbackquote}\ T} is the least set @@ -73,7 +73,7 @@ First we prove monotonicity of the function inside \isa{lfp}% \end{isamarkuptext}% -\isacommand{lemma}\ mono{\isacharunderscore}ef{\isacharcolon}\ {\isachardoublequote}mono{\isacharparenleft}{\isasymlambda}T{\isachardot}\ A\ {\isasymunion}\ M{\isacharcircum}{\isacharminus}{\isadigit{1}}\ {\isacharbackquote}{\isacharbackquote}{\isacharbackquote}\ T{\isacharparenright}{\isachardoublequote}\isanewline +\isacommand{lemma}\ mono{\isacharunderscore}ef{\isacharcolon}\ {\isachardoublequote}mono{\isacharparenleft}{\isasymlambda}T{\isachardot}\ A\ {\isasymunion}\ M{\isasyminverse}\ {\isacharbackquote}{\isacharbackquote}{\isacharbackquote}\ T{\isacharparenright}{\isachardoublequote}\isanewline \isacommand{apply}{\isacharparenleft}rule\ monoI{\isacharparenright}\isanewline \isacommand{apply}\ blast\isanewline \isacommand{done}% @@ -85,7 +85,7 @@ a separate lemma:% \end{isamarkuptext}% \isacommand{lemma}\ EF{\isacharunderscore}lemma{\isacharcolon}\isanewline -\ \ {\isachardoublequote}lfp{\isacharparenleft}{\isasymlambda}T{\isachardot}\ A\ {\isasymunion}\ M{\isacharcircum}{\isacharminus}{\isadigit{1}}\ {\isacharbackquote}{\isacharbackquote}{\isacharbackquote}\ T{\isacharparenright}\ {\isacharequal}\ {\isacharbraceleft}s{\isachardot}\ {\isasymexists}t{\isachardot}\ {\isacharparenleft}s{\isacharcomma}t{\isacharparenright}\ {\isasymin}\ M{\isacharcircum}{\isacharasterisk}\ {\isasymand}\ t\ {\isasymin}\ A{\isacharbraceright}{\isachardoublequote}% +\ \ {\isachardoublequote}lfp{\isacharparenleft}{\isasymlambda}T{\isachardot}\ A\ {\isasymunion}\ M{\isasyminverse}\ {\isacharbackquote}{\isacharbackquote}{\isacharbackquote}\ T{\isacharparenright}\ {\isacharequal}\ {\isacharbraceleft}s{\isachardot}\ {\isasymexists}t{\isachardot}\ {\isacharparenleft}s{\isacharcomma}t{\isacharparenright}\ {\isasymin}\ M\isactrlsup {\isacharasterisk}\ {\isasymand}\ t\ {\isasymin}\ A{\isacharbraceright}{\isachardoublequote}% \begin{isamarkuptxt}% \noindent The equality is proved in the canonical fashion by proving that each set @@ -110,11 +110,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 \isa{blast} with the help of transitivity of \isa{{\isacharcircum}{\isacharasterisk}}:% +which is proved by \isa{blast} with the help of transitivity of \isa{\isactrlsup {\isacharasterisk}}:% \end{isamarkuptxt}% \ \isacommand{apply}{\isacharparenleft}blast\ intro{\isacharcolon}\ rtrancl{\isacharunderscore}trans{\isacharparenright}% \begin{isamarkuptxt}% diff -r 2d4c058749a7 -r c00ac928fc6f doc-src/TutorialI/appendix.tex --- a/doc-src/TutorialI/appendix.tex Sat Jan 06 11:27:09 2001 +0100 +++ b/doc-src/TutorialI/appendix.tex Sat Jan 06 12:39:05 2001 +0100 @@ -89,6 +89,12 @@ \isasymInter\index{$HOL3Set2@\isasymInter|bold}& \ttindexbold{INT}, \ttindexbold{Inter} & \verb$\$\\ +\isactrlsup{\isacharasterisk}\index{$HOL4star@\isactrlsup{\isacharasterisk}|bold}& +\verb$^*$\index{$HOL4star@\verb$^$\texttt{*}|bold} & +\verb$\<^sup>*$\\ +\isasyminverse\index{$HOL4inv@\isasyminverse|bold}& +\verb$^-1$\index{$HOL4inv@\verb$^-1$|bold} & +\verb$\$\\ \hline \end{tabular} \end{center}