diff -r a0460a450cf9 -r bbfc360db011 doc-src/TutorialI/Overview/LNCS/Sets.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/TutorialI/Overview/LNCS/Sets.thy Mon Jul 01 15:33:03 2002 +0200 @@ -0,0 +1,143 @@ +(*<*)theory Sets = Main:(*>*) + +section{*Sets, Functions and Relations*} + +subsection{*Set Notation*} + +text{* +\begin{center} +\begin{tabular}{ccc} +@{term "A \ B"} & @{term "A \ B"} & @{term "A - B"} \\ +@{term "a \ A"} & @{term "b \ A"} \\ +@{term "{a,b}"} & @{text "{x. P x}"} \\ +@{term "\ M"} & @{text "\a \ A. F a"} +\end{tabular} +\end{center}*} +(*<*)term "A \ B" term "A \ B" term "A - B" +term "a \ A" term "b \ A" +term "{a,b}" term "{x. P x}" +term "\ M" term "\a \ A. F a"(*>*) + +subsection{*Some Functions*} + +text{* +\begin{tabular}{l} +@{thm id_def}\\ +@{thm o_def[no_vars]}\\ +@{thm image_def[no_vars]}\\ +@{thm vimage_def[no_vars]} +\end{tabular}*} +(*<*)thm id_def o_def[no_vars] image_def[no_vars] vimage_def[no_vars](*>*) + +subsection{*Some Relations*} + +text{* +\begin{tabular}{l} +@{thm Id_def}\\ +@{thm converse_def[no_vars]}\\ +@{thm Image_def[no_vars]}\\ +@{thm rtrancl_refl[no_vars]}\\ +@{thm rtrancl_into_rtrancl[no_vars]}\\ +@{thm trancl_def[no_vars]} +\end{tabular}*} +(*<*)thm Id_def +thm converse_def[no_vars] +thm Image_def[no_vars] +thm relpow.simps[no_vars] +thm rtrancl.intros[no_vars] +thm trancl_def[no_vars](*>*) + +subsection{*Wellfoundedness*} + +text{* +\begin{tabular}{l} +@{thm wf_def[no_vars]}\\ +@{thm wf_iff_no_infinite_down_chain[no_vars]} +\end{tabular}*} +(*<*)thm wf_def[no_vars] +thm wf_iff_no_infinite_down_chain[no_vars](*>*) + +subsection{*Fixed Point Operators*} + +text{* +\begin{tabular}{l} +@{thm lfp_def[no_vars]}\\ +@{thm lfp_unfold[no_vars]}\\ +@{thm lfp_induct[no_vars]} +\end{tabular}*} +(*<*)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 clarsimp +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(*>*)