diff -r 6902638af59e -r 860c65c7388a doc-src/TutorialI/Overview/Sets.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/TutorialI/Overview/Sets.thy Fri Mar 30 16:12:57 2001 +0200 @@ -0,0 +1,123 @@ +theory Sets = Main: + +section{*Sets, Functions and Relations*} + +subsection{*Set Notation*} + +term "A \ B" +term "A \ B" +term "A - B" +term "a \ A" +term "b \ A" +term "{a,b}" +term "{x. P x}" +term "{x+y+eps |x y. x < y}" +term "\ M" +term "\a \ A. F a" + +subsection{*Functions*} + +thm id_def +thm o_assoc +thm image_Int +thm vimage_Compl + + +subsection{*Relations*} + +thm Id_def +thm converse_comp +thm Image_def +thm relpow.simps +thm rtrancl_idemp +thm trancl_converse + +subsection{*Wellfoundedness*} + +thm wf_def +thm wf_iff_no_infinite_down_chain + + +subsection{*Fixed Point Operators*} + +thm lfp_def gfp_def +thm lfp_unfold +thm lfp_induct + + +subsection{*Case Study: Verified Model Checking*} + + +typedecl state + +consts M :: "(state \ state)set"; + +typedecl atom + +consts L :: "state \ atom set" + +datatype formula = Atom atom + | Neg formula + | And formula formula + | AX formula + | EF formula + +consts valid :: "state \ formula \ bool" ("(_ \ _)" [80,80] 80) + +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)" +"s \ EF f = (\t. (s,t) \ M\<^sup>* \ t \ f)"; + +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}" +"mc(EF f) = lfp(\T. mc f \ (M\ `` T))" + +lemma mono_ef: "mono(\T. A \ (M\ `` T))" +apply(rule monoI) +apply blast +done + +lemma EF_lemma: + "lfp(\T. A \ (M\ `` T)) = {s. \t. (s,t) \ M\<^sup>* \ t \ A}" +apply(rule equalityI) + thm lfp_lowerbound + apply(rule lfp_lowerbound) + apply(blast intro: rtrancl_trans); +apply(rule subsetI) +apply(simp, clarify) +apply(erule converse_rtrancl_induct) +thm lfp_unfold[OF mono_ef] + apply(subst lfp_unfold[OF mono_ef]) + apply(blast) +apply(subst lfp_unfold[OF mono_ef]) +apply(blast) +done + +theorem "mc f = {s. s \ f}"; +apply(induct_tac f); +apply(auto simp add:EF_lemma); +done; + +text{* +\begin{exercise} +@{term AX} has a dual operator @{term EN}\footnote{We cannot use the customary @{text EX} +as that is the \textsc{ascii}-equivalent of @{text"\"}} +(``there exists a next state such that'') with the intended semantics +@{prop[display]"(s \ EN f) = (EX t. (s,t) : M & t \ f)"} +Fortunately, @{term"EN f"} can already be expressed as a PDL formula. How? + +Show that the semantics for @{term EF} satisfies the following recursion equation: +@{prop[display]"(s \ EF f) = (s \ f | s \ EN(EF f))"} +\end{exercise} +*} + +end + +