# HG changeset patch # User wenzelm # Date 1451142600 -3600 # Node ID 02610a8064678251e5afc979d478dc6e2ee28290 # Parent cf58b5b794b2156a1ce0e34a895e3fdc254d652a tuned; diff -r cf58b5b794b2 -r 02610a806467 src/HOL/ex/CTL.thy --- a/src/HOL/ex/CTL.thy Sat Dec 26 15:59:27 2015 +0100 +++ b/src/HOL/ex/CTL.thy Sat Dec 26 16:10:00 2015 +0100 @@ -9,16 +9,16 @@ begin text \ - We formalize basic concepts of Computational Tree Logic (CTL) - @{cite "McMillan-PhDThesis" and "McMillan-LectureNotes"} within the - simply-typed set theory of HOL. + We formalize basic concepts of Computational Tree Logic (CTL) @{cite + "McMillan-PhDThesis" and "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. + 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 @@ -34,55 +34,52 @@ text \ - \smallskip The CTL path operators are more interesting; they are - based on an arbitrary, but fixed model \\\, which is simply - a transition relation over states @{typ "'a"}. + \<^smallskip> + The CTL path operators are more interesting; they are based on an arbitrary, + but fixed model \\\, which is simply a transition relation over states + @{typ 'a}. \ axiomatization \ :: "('a \ 'a) set" text \ - The operators \\\, \\\, \\\ are taken - as primitives, while \\\, \\\, \\\ are - defined as derived ones. The formula \\ 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 \\ p\ holds in a state @{term - s}, iff there is a path in \\\, 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 \\ 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 \\ p\ and \\ p\ may be expressed using least and greatest fixed points - @{cite "McMillan-PhDThesis"}. + The operators \\\, \\\, \\\ are taken as primitives, while + \\\, \\\, \\\ are defined as derived ones. The formula \\ 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 \\ p\ holds in a state @{term s}, iff there is a path in + \\\, 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 \\ 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 \\ p\ and \\ p\ may be expressed using + least and greatest fixed points @{cite "McMillan-PhDThesis"}. \ -definition - EX ("\ _" [80] 90) where "\ p = {s. \s'. (s, s') \ \ \ s' \ p}" -definition - EF ("\ _" [80] 90) where "\ p = lfp (\s. p \ \ s)" -definition - EG ("\ _" [80] 90) where "\ p = gfp (\s. p \ \ s)" +definition EX ("\ _" [80] 90) + where [simp]: "\ p = {s. \s'. (s, s') \ \ \ s' \ p}" +definition EF ("\ _" [80] 90) + where [simp]: "\ p = lfp (\s. p \ \ s)" +definition EG ("\ _" [80] 90) + where [simp]: "\ p = gfp (\s. p \ \ s)" text \ - \\\, \\\ and \\\ are now defined - dually in terms of \\\, \\\ and \\\. + \\\, \\\ and \\\ are now defined dually in terms of \\\, + \\\ and \\\. \ -definition - AX ("\ _" [80] 90) where "\ p = - \ - p" -definition - AF ("\ _" [80] 90) where "\ p = - \ - p" -definition - AG ("\ _" [80] 90) where "\ p = - \ - p" - -lemmas [simp] = EX_def EG_def AX_def EF_def AF_def AG_def +definition AX ("\ _" [80] 90) + where [simp]: "\ p = - \ - p" +definition AF ("\ _" [80] 90) + where [simp]: "\ p = - \ - p" +definition AG ("\ _" [80] 90) + where [simp]: "\ p = - \ - p" subsection \Basic fixed point properties\ text \ - First of all, we use the de-Morgan property of fixed points + First of all, we use the de-Morgan property of fixed points. \ lemma lfp_gfp: "lfp f = - gfp (\s::'a set. - (f (- s)))" @@ -118,12 +115,14 @@ by (simp add: lfp_gfp) text \ - in order to give dual fixed point representations of @{term "AF p"} - and @{term "AG p"}: + In order to give dual fixed point representations of @{term "\ p"} and + @{term "\ 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 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 - @@ -144,10 +143,10 @@ 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"}. + 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" @@ -157,9 +156,8 @@ qed text \ - This fact may be split up into two inequalities (merely using - transitivity of \\\, which is an instance of the overloaded - \\\ in Isabelle/HOL). + This fact may be split up into two inequalities (merely using transitivity + of \\\, which is an instance of the overloaded \\\ in Isabelle/HOL). \ lemma AG_fp_1: "\ p \ p" @@ -175,11 +173,11 @@ 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"}. + On the other hand, we have from the Knaster-Tarski fixed point theorem that + any other post-fixed point of @{term "\s. p \ \ s"} is smaller than + @{term "\ p"}. A post-fixed point is a set of states @{term q} such that + @{term "q \ p \ \ q"}. This leads to the following co-induction + principle for @{term "\ p"}. \ lemma AG_I: "q \ p \ \ q \ q \ \ p" @@ -189,10 +187,10 @@ subsection \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 \AG\ (see below). We will use some elementary - monotonicity and distributivity rules. + 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 \\\ + (see below). We will use some elementary monotonicity and distributivity + rules. \ lemma AX_int: "\ (p \ q) = \ p \ \ q" by auto @@ -201,21 +199,21 @@ by (simp only: AG_gfp, rule gfp_mono) auto text \ - The formula @{term "AG p"} implies @{term "AX p"} (we use - substitution of \\\ with monotonicity). + The formula @{term "AG p"} implies @{term "AX p"} (we use substitution of + \\\ 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 + also have "\ p \ p" by (rule AG_fp_1) + moreover note AX_mono finally show ?thesis . qed text \ - Furthermore we show idempotency of the \\\ operator. - The proof is a good example of how accumulated facts may get - used to feed a single rule step. + Furthermore we show idempotency of the \\\ 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" @@ -231,13 +229,14 @@ qed text \ - \smallskip We now give an alternative characterization of the \\\ operator, which describes the \\\ 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. + \<^smallskip> + We now give an alternative characterization of the \\\ operator, which + describes the \\\ 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" @@ -284,9 +283,8 @@ subsection \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 - \\\ and \\\ commute. + Further interesting properties of CTL expressions may be demonstrated with + the help of tree induction; here we show that \\\ and \\\ commute. \ theorem AG_AX_commute: "\ \ p = \ \ p"