diff -r 78822f2d921f -r 67f2920862c7 doc-src/TutorialI/CTL/PDL.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/TutorialI/CTL/PDL.thy Thu Sep 14 17:46:00 2000 +0200 @@ -0,0 +1,58 @@ +theory PDL = Main:; + +typedecl atom; +types state = "atom set"; + +datatype ctl_form = Atom atom + | NOT ctl_form + | And ctl_form ctl_form + | AX ctl_form + | EF ctl_form; + +consts valid :: "state \ ctl_form \ bool" ("(_ \ _)" [80,80] 80) + M :: "(state \ state)set"; + +primrec +"s \ Atom a = (a\s)" +"s \ NOT 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)"; + +consts mc :: "ctl_form \ state set"; +primrec +"mc(Atom a) = {s. a\s}" +"mc(NOT f) = -mc f" +"mc(And f g) = mc f Int mc g" +"mc(AX f) = {s. \t. (s,t) \ M \ t \ mc f}" +"mc(EF f) = lfp(\T. mc f \ {s. \t. (s,t)\M \ t\T})"; + +lemma mono_lemma: "mono(\T. A \ {s. \t. (s,t)\M \ t\T})"; +apply(rule monoI); +by(blast); + +lemma lfp_conv_EF: +"lfp(\T. A \ {s. \t. (s,t)\M \ t\T}) = {s. \t. (s,t) \ M^* \ t \ A}"; +apply(rule equalityI); + apply(rule subsetI); + apply(simp); + apply(erule Lfp.induct); + apply(rule mono_lemma); + apply(simp); + apply(blast intro: r_into_rtrancl rtrancl_trans); +apply(rule subsetI); +apply(simp); +apply(erule exE); +apply(erule conjE); +apply(erule_tac P = "t\A" in rev_mp); +apply(erule converse_rtrancl_induct); + apply(rule ssubst[OF lfp_Tarski[OF mono_lemma]]); + apply(blast); +apply(rule ssubst[OF lfp_Tarski[OF mono_lemma]]); +by(blast); + +theorem "mc f = {s. s \ f}"; +apply(induct_tac f); +by(auto simp add:lfp_conv_EF); + +end;