diff -r 4320bce5873f -r e524119dbf19 src/HOL/ex/CTL.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/ex/CTL.thy Thu Apr 28 17:08:08 2005 +0200 @@ -0,0 +1,317 @@ + +(* Title: HOL/ex/CTL.thy + ID: $Id$ + Author: Gertrud Bauer +*) + +header {* CTL formulae *} + +theory CTL = Main: + + + +text {* + We formalize basic concepts of Computational Tree Logic (CTL) + \cite{McMillan-PhDThesis,McMillan-LectureNotes} within the + simply-typed set theory of HOL. + + By using the common technique of ``shallow embedding'', a CTL + formula is identified with the corresponding set of states where it + holds. Consequently, CTL operations such as negation, conjunction, + disjunction simply become complement, intersection, union of sets. + We only require a separate operation for implication, as point-wise + inclusion is usually not encountered in plain set-theory. +*} + +lemmas [intro!] = Int_greatest Un_upper2 Un_upper1 Int_lower1 Int_lower2 + +types 'a ctl = "'a set" +constdefs + imp :: "'a ctl \ 'a ctl \ 'a ctl" (infixr "\" 75) + "p \ q \ - p \ q" + +lemma [intro!]: "p \ p \ q \ q" by (unfold imp_def) auto +lemma [intro!]: "p \ (q \ p)" by (unfold imp_def) rule + + +text {* + \smallskip The CTL path operators are more interesting; they are + based on an arbitrary, but fixed model @{text \}, which is simply + a transition relation over states @{typ "'a"}. +*} + +consts model :: "('a \ 'a) set" ("\") + +text {* + The operators @{text \}, @{text \}, @{text \} are taken + as primitives, while @{text \}, @{text \}, @{text \} are + defined as derived ones. The formula @{text "\ p"} holds in a + state @{term s}, iff there is a successor state @{term s'} (with + respect to the model @{term \}), such that @{term p} holds in + @{term s'}. The formula @{text "\ p"} holds in a state @{term + s}, iff there is a path in @{text \}, starting from @{term s}, + such that there exists a state @{term s'} on the path, such that + @{term p} holds in @{term s'}. The formula @{text "\ p"} holds + in a state @{term s}, iff there is a path, starting from @{term s}, + such that for all states @{term s'} on the path, @{term p} holds in + @{term s'}. It is easy to see that @{text "\ p"} and @{text + "\ p"} may be expressed using least and greatest fixed points + \cite{McMillan-PhDThesis}. +*} + +constdefs + EX :: "'a ctl \ 'a ctl" ("\ _" [80] 90) "\ p \ {s. \s'. (s, s') \ \ \ s' \ p}" + EF :: "'a ctl \ 'a ctl" ("\ _" [80] 90) "\ p \ lfp (\s. p \ \ s)" + EG :: "'a ctl \ 'a ctl" ("\ _" [80] 90) "\ p \ gfp (\s. p \ \ s)" + +text {* + @{text "\"}, @{text "\"} and @{text "\"} are now defined + dually in terms of @{text "\"}, @{text "\"} and @{text + "\"}. +*} + +constdefs + AX :: "'a ctl \ 'a ctl" ("\ _" [80] 90) "\ p \ - \ - p" + AF :: "'a ctl \ 'a ctl" ("\ _" [80] 90) "\ p \ - \ - p" + AG :: "'a ctl \ 'a ctl" ("\ _" [80] 90) "\ p \ - \ - p" + +lemmas [simp] = EX_def EG_def AX_def EF_def AF_def AG_def + + +section {* Basic fixed point properties *} + +text {* + First of all, we use the de-Morgan property of fixed points +*} + +lemma lfp_gfp: "lfp f = - gfp (\s . - (f (- s)))" +proof + show "lfp f \ - gfp (\s. - f (- s))" + proof + fix x assume l: "x \ lfp f" + show "x \ - gfp (\s. - f (- s))" + proof + assume "x \ gfp (\s. - f (- s))" + then obtain u where "x \ u" and "u \ - f (- u)" by (unfold gfp_def) auto + then have "f (- u) \ - u" by auto + then have "lfp f \ - u" by (rule lfp_lowerbound) + from l and this have "x \ u" by auto + then show False by contradiction + qed + qed + show "- gfp (\s. - f (- s)) \ lfp f" + proof (rule lfp_greatest) + fix u assume "f u \ u" + then have "- u \ - f u" by auto + then have "- u \ - f (- (- u))" by simp + then have "- u \ gfp (\s. - f (- s))" by (rule gfp_upperbound) + then show "- gfp (\s. - f (- s)) \ u" by auto + qed +qed + +lemma lfp_gfp': "- lfp f = gfp (\s. - (f (- s)))" + by (simp add: lfp_gfp) + +lemma gfp_lfp': "- gfp f = lfp (\s. - (f (- s)))" + by (simp add: lfp_gfp) + +text {* + in order to give dual fixed point representations of @{term "AF p"} + and @{term "AG p"}: +*} + +lemma AF_lfp: "\ p = lfp (\s. p \ \ s)" by (simp add: lfp_gfp) +lemma AG_gfp: "\ p = gfp (\s. p \ \ s)" by (simp add: lfp_gfp) + +lemma EF_fp: "\ p = p \ \ \ p" +proof - + have "mono (\s. p \ \ s)" by rule (auto simp add: EX_def) + then show ?thesis by (simp only: EF_def) (rule lfp_unfold) +qed + +lemma AF_fp: "\ p = p \ \ \ p" +proof - + have "mono (\s. p \ \ s)" by rule (auto simp add: AX_def EX_def) + then show ?thesis by (simp only: AF_lfp) (rule lfp_unfold) +qed + +lemma EG_fp: "\ p = p \ \ \ p" +proof - + have "mono (\s. p \ \ s)" by rule (auto simp add: EX_def) + then show ?thesis by (simp only: EG_def) (rule gfp_unfold) +qed + +text {* + From the greatest fixed point definition of @{term "\ p"}, we + derive as a consequence of the Knaster-Tarski theorem on the one + hand that @{term "\ p"} is a fixed point of the monotonic + function @{term "\s. p \ \ s"}. +*} + +lemma AG_fp: "\ p = p \ \ \ p" +proof - + have "mono (\s. p \ \ s)" by rule (auto simp add: AX_def EX_def) + then show ?thesis by (simp only: AG_gfp) (rule gfp_unfold) +qed + +text {* + This fact may be split up into two inequalities (merely using + transitivity of @{text "\" }, which is an instance of the overloaded + @{text "\"} in Isabelle/HOL). +*} + +lemma AG_fp_1: "\ p \ p" +proof - + note AG_fp also have "p \ \ \ p \ p" by auto + finally show ?thesis . +qed + +text {**} + +lemma AG_fp_2: "\ p \ \ \ p" +proof - + note AG_fp also have "p \ \ \ p \ \ \ p" by auto + finally show ?thesis . +qed + +text {* + On the other hand, we have from the Knaster-Tarski fixed point + theorem that any other post-fixed point of @{term "\s. p \ AX s"} is + smaller than @{term "AG p"}. A post-fixed point is a set of states + @{term q} such that @{term "q \ p \ AX q"}. This leads to the + following co-induction principle for @{term "AG p"}. +*} + +lemma AG_I: "q \ p \ \ q \ q \ \ p" + by (simp only: AG_gfp) (rule gfp_upperbound) + + +section {* The tree induction principle \label{sec:calc-ctl-tree-induct} *} + +text {* + With the most basic facts available, we are now able to establish a + few more interesting results, leading to the \emph{tree induction} + principle for @{text AG} (see below). We will use some elementary + monotonicity and distributivity rules. +*} + +lemma AX_int: "\ (p \ q) = \ p \ \ q" by auto +lemma AX_mono: "p \ q \ \ p \ \ q" by auto +lemma AG_mono: "p \ q \ \ p \ \ q" + by (simp only: AG_gfp, rule gfp_mono) auto + +text {* + The formula @{term "AG p"} implies @{term "AX p"} (we use + substitution of @{text "\"} with monotonicity). +*} + +lemma AG_AX: "\ p \ \ p" +proof - + have "\ p \ \ \ p" by (rule AG_fp_2) + also have "\ p \ p" by (rule AG_fp_1) moreover note AX_mono + finally show ?thesis . +qed + +text {* + Furthermore we show idempotency of the @{text "\"} operator. + The proof is a good example of how accumulated facts may get + used to feed a single rule step. +*} + +lemma AG_AG: "\ \ p = \ p" +proof + show "\ \ p \ \ p" by (rule AG_fp_1) +next + show "\ p \ \ \ p" + proof (rule AG_I) + have "\ p \ \ p" .. + moreover have "\ p \ \ \ p" by (rule AG_fp_2) + ultimately show "\ p \ \ p \ \ \ p" .. + qed +qed + +text {* + \smallskip We now give an alternative characterization of the @{text + "\"} operator, which describes the @{text "\"} operator in + an ``operational'' way by tree induction: In a state holds @{term + "AG p"} iff in that state holds @{term p}, and in all reachable + states @{term s} follows from the fact that @{term p} holds in + @{term s}, that @{term p} also holds in all successor states of + @{term s}. We use the co-induction principle @{thm [source] AG_I} + to establish this in a purely algebraic manner. +*} + +theorem AG_induct: "p \ \ (p \ \ p) = \ p" +proof + show "p \ \ (p \ \ p) \ \ p" (is "?lhs \ _") + proof (rule AG_I) + show "?lhs \ p \ \ ?lhs" + proof + show "?lhs \ p" .. + show "?lhs \ \ ?lhs" + proof - + { + have "\ (p \ \ p) \ p \ \ p" by (rule AG_fp_1) + also have "p \ p \ \ p \ \ p" .. + finally have "?lhs \ \ p" by auto + } + moreover + { + have "p \ \ (p \ \ p) \ \ (p \ \ p)" .. + also have "\ \ \ \" by (rule AG_fp_2) + finally have "?lhs \ \ \ (p \ \ p)" . + } + ultimately have "?lhs \ \ p \ \ \ (p \ \ p)" .. + also have "\ = \ ?lhs" by (simp only: AX_int) + finally show ?thesis . + qed + qed + qed +next + show "\ p \ p \ \ (p \ \ p)" + proof + show "\ p \ p" by (rule AG_fp_1) + show "\ p \ \ (p \ \ p)" + proof - + have "\ p = \ \ p" by (simp only: AG_AG) + also have "\ p \ \ p" by (rule AG_AX) moreover note AG_mono + also have "\ p \ (p \ \ p)" .. moreover note AG_mono + finally show ?thesis . + qed + qed +qed + + +section {* An application of tree induction \label{sec:calc-ctl-commute} *} + +text {* + Further interesting properties of CTL expressions may be + demonstrated with the help of tree induction; here we show that + @{text \} and @{text \} commute. +*} + +theorem AG_AX_commute: "\ \ p = \ \ p" +proof - + have "\ \ p = \ p \ \ \ \ p" by (rule AG_fp) + also have "\ = \ (p \ \ \ p)" by (simp only: AX_int) + also have "p \ \ \ p = \ p" (is "?lhs = _") + proof + have "\ p \ p \ \ p" .. + also have "p \ \ (p \ \ p) = \ p" by (rule AG_induct) + also note Int_mono AG_mono + ultimately show "?lhs \ \ p" by fast + next + have "\ p \ p" by (rule AG_fp_1) + moreover + { + have "\ p = \ \ p" by (simp only: AG_AG) + also have "\ p \ \ p" by (rule AG_AX) + also note AG_mono + ultimately have "\ p \ \ \ p" . + } + ultimately show "\ p \ ?lhs" .. + qed + finally show ?thesis . +qed + +end