| author | Fabian Huch <huch@in.tum.de> | 
| Mon, 10 Jun 2024 14:08:15 +0200 | |
| changeset 80335 | b835b40f53ec | 
| parent 77036 | d0151eb9ecb0 | 
| child 80914 | d97fdabd9e2b | 
| permissions | -rw-r--r-- | 
| 15871 | 1 | (* Title: HOL/ex/CTL.thy | 
| 2 | Author: Gertrud Bauer | |
| 3 | *) | |
| 4 | ||
| 61343 | 5 | section \<open>CTL formulae\<close> | 
| 15871 | 6 | |
| 46685 | 7 | theory CTL | 
| 8 | imports Main | |
| 9 | begin | |
| 15871 | 10 | |
| 61343 | 11 | text \<open> | 
| 77036 | 12 | We formalize basic concepts of Computational Tree Logic (CTL) \<^cite>\<open>"McMillan-PhDThesis"\<close> within the simply-typed | 
| 61934 | 13 | set theory of HOL. | 
| 15871 | 14 | |
| 61934 | 15 | By using the common technique of ``shallow embedding'', a CTL formula is | 
| 16 | identified with the corresponding set of states where it holds. | |
| 17 | Consequently, CTL operations such as negation, conjunction, disjunction | |
| 18 | simply become complement, intersection, union of sets. We only require a | |
| 19 | separate operation for implication, as point-wise inclusion is usually not | |
| 20 | encountered in plain set-theory. | |
| 61343 | 21 | \<close> | 
| 15871 | 22 | |
| 23 | lemmas [intro!] = Int_greatest Un_upper2 Un_upper1 Int_lower1 Int_lower2 | |
| 24 | ||
| 42463 | 25 | type_synonym 'a ctl = "'a set" | 
| 20807 | 26 | |
| 63054 | 27 | definition imp :: "'a ctl \<Rightarrow> 'a ctl \<Rightarrow> 'a ctl" (infixr "\<rightarrow>" 75) | 
| 28 | where "p \<rightarrow> q = - p \<union> q" | |
| 15871 | 29 | |
| 20807 | 30 | lemma [intro!]: "p \<inter> p \<rightarrow> q \<subseteq> q" unfolding imp_def by auto | 
| 31 | lemma [intro!]: "p \<subseteq> (q \<rightarrow> p)" unfolding imp_def by rule | |
| 15871 | 32 | |
| 33 | ||
| 61343 | 34 | text \<open> | 
| 61934 | 35 | \<^smallskip> | 
| 36 | The CTL path operators are more interesting; they are based on an arbitrary, | |
| 37 | but fixed model \<open>\<M>\<close>, which is simply a transition relation over states | |
| 69597 | 38 | \<^typ>\<open>'a\<close>. | 
| 61343 | 39 | \<close> | 
| 15871 | 40 | |
| 20807 | 41 | axiomatization \<M> :: "('a \<times> 'a) set"
 | 
| 15871 | 42 | |
| 61343 | 43 | text \<open> | 
| 63055 | 44 | The operators \<open>\<^bold>E\<^bold>X\<close>, \<open>\<^bold>E\<^bold>F\<close>, \<open>\<^bold>E\<^bold>G\<close> are taken as primitives, while \<open>\<^bold>A\<^bold>X\<close>, | 
| 45 | \<open>\<^bold>A\<^bold>F\<close>, \<open>\<^bold>A\<^bold>G\<close> are defined as derived ones. The formula \<open>\<^bold>E\<^bold>X p\<close> holds in a | |
| 46 | state \<open>s\<close>, iff there is a successor state \<open>s'\<close> (with respect to the model | |
| 47 | \<open>\<M>\<close>), such that \<open>p\<close> holds in \<open>s'\<close>. The formula \<open>\<^bold>E\<^bold>F p\<close> holds in a state | |
| 48 | \<open>s\<close>, iff there is a path in \<open>\<M>\<close>, starting from \<open>s\<close>, such that there exists a | |
| 49 | state \<open>s'\<close> on the path, such that \<open>p\<close> holds in \<open>s'\<close>. The formula \<open>\<^bold>E\<^bold>G p\<close> | |
| 50 | holds in a state \<open>s\<close>, iff there is a path, starting from \<open>s\<close>, such that for | |
| 51 | all states \<open>s'\<close> on the path, \<open>p\<close> holds in \<open>s'\<close>. It is easy to see that \<open>\<^bold>E\<^bold>F | |
| 52 | p\<close> and \<open>\<^bold>E\<^bold>G p\<close> may be expressed using least and greatest fixed points | |
| 76987 | 53 | \<^cite>\<open>"McMillan-PhDThesis"\<close>. | 
| 61343 | 54 | \<close> | 
| 15871 | 55 | |
| 63055 | 56 | definition EX  ("\<^bold>E\<^bold>X _" [80] 90)
 | 
| 57 |   where [simp]: "\<^bold>E\<^bold>X p = {s. \<exists>s'. (s, s') \<in> \<M> \<and> s' \<in> p}"
 | |
| 58 | ||
| 59 | definition EF ("\<^bold>E\<^bold>F _" [80] 90)
 | |
| 60 | where [simp]: "\<^bold>E\<^bold>F p = lfp (\<lambda>s. p \<union> \<^bold>E\<^bold>X s)" | |
| 61 | ||
| 62 | definition EG ("\<^bold>E\<^bold>G _" [80] 90)
 | |
| 63 | where [simp]: "\<^bold>E\<^bold>G p = gfp (\<lambda>s. p \<inter> \<^bold>E\<^bold>X s)" | |
| 15871 | 64 | |
| 61343 | 65 | text \<open> | 
| 63055 | 66 | \<open>\<^bold>A\<^bold>X\<close>, \<open>\<^bold>A\<^bold>F\<close> and \<open>\<^bold>A\<^bold>G\<close> are now defined dually in terms of \<open>\<^bold>E\<^bold>X\<close>, | 
| 67 | \<open>\<^bold>E\<^bold>F\<close> and \<open>\<^bold>E\<^bold>G\<close>. | |
| 61343 | 68 | \<close> | 
| 15871 | 69 | |
| 63055 | 70 | definition AX  ("\<^bold>A\<^bold>X _" [80] 90)
 | 
| 71 | where [simp]: "\<^bold>A\<^bold>X p = - \<^bold>E\<^bold>X - p" | |
| 72 | definition AF  ("\<^bold>A\<^bold>F _" [80] 90)
 | |
| 73 | where [simp]: "\<^bold>A\<^bold>F p = - \<^bold>E\<^bold>G - p" | |
| 74 | definition AG  ("\<^bold>A\<^bold>G _" [80] 90)
 | |
| 75 | where [simp]: "\<^bold>A\<^bold>G p = - \<^bold>E\<^bold>F - p" | |
| 15871 | 76 | |
| 77 | ||
| 61343 | 78 | subsection \<open>Basic fixed point properties\<close> | 
| 15871 | 79 | |
| 61343 | 80 | text \<open> | 
| 61934 | 81 | First of all, we use the de-Morgan property of fixed points. | 
| 61343 | 82 | \<close> | 
| 15871 | 83 | |
| 21026 | 84 | lemma lfp_gfp: "lfp f = - gfp (\<lambda>s::'a set. - (f (- s)))" | 
| 15871 | 85 | proof | 
| 86 | show "lfp f \<subseteq> - gfp (\<lambda>s. - f (- s))" | |
| 87 | proof | |
| 63054 | 88 | show "x \<in> - gfp (\<lambda>s. - f (- s))" if l: "x \<in> lfp f" for x | 
| 15871 | 89 | proof | 
| 90 | assume "x \<in> gfp (\<lambda>s. - f (- s))" | |
| 21026 | 91 | then obtain u where "x \<in> u" and "u \<subseteq> - f (- u)" | 
| 32587 
caa5ada96a00
Inter and Union are mere abbreviations for Inf and Sup
 haftmann parents: 
26813diff
changeset | 92 | by (auto simp add: gfp_def) | 
| 15871 | 93 | then have "f (- u) \<subseteq> - u" by auto | 
| 94 | then have "lfp f \<subseteq> - u" by (rule lfp_lowerbound) | |
| 95 | from l and this have "x \<notin> u" by auto | |
| 61343 | 96 | with \<open>x \<in> u\<close> show False by contradiction | 
| 15871 | 97 | qed | 
| 98 | qed | |
| 99 | show "- gfp (\<lambda>s. - f (- s)) \<subseteq> lfp f" | |
| 100 | proof (rule lfp_greatest) | |
| 63054 | 101 | fix u | 
| 102 | assume "f u \<subseteq> u" | |
| 15871 | 103 | then have "- u \<subseteq> - f u" by auto | 
| 104 | then have "- u \<subseteq> - f (- (- u))" by simp | |
| 105 | then have "- u \<subseteq> gfp (\<lambda>s. - f (- s))" by (rule gfp_upperbound) | |
| 106 | then show "- gfp (\<lambda>s. - f (- s)) \<subseteq> u" by auto | |
| 107 | qed | |
| 108 | qed | |
| 109 | ||
| 21026 | 110 | lemma lfp_gfp': "- lfp f = gfp (\<lambda>s::'a set. - (f (- s)))" | 
| 15871 | 111 | by (simp add: lfp_gfp) | 
| 112 | ||
| 21026 | 113 | lemma gfp_lfp': "- gfp f = lfp (\<lambda>s::'a set. - (f (- s)))" | 
| 15871 | 114 | by (simp add: lfp_gfp) | 
| 115 | ||
| 61343 | 116 | text \<open> | 
| 69597 | 117 | In order to give dual fixed point representations of \<^term>\<open>\<^bold>A\<^bold>F p\<close> and | 
| 118 | \<^term>\<open>\<^bold>A\<^bold>G p\<close>: | |
| 61343 | 119 | \<close> | 
| 15871 | 120 | |
| 63055 | 121 | lemma AF_lfp: "\<^bold>A\<^bold>F p = lfp (\<lambda>s. p \<union> \<^bold>A\<^bold>X s)" | 
| 61934 | 122 | by (simp add: lfp_gfp) | 
| 15871 | 123 | |
| 63055 | 124 | lemma AG_gfp: "\<^bold>A\<^bold>G p = gfp (\<lambda>s. p \<inter> \<^bold>A\<^bold>X s)" | 
| 125 | by (simp add: lfp_gfp) | |
| 126 | ||
| 127 | lemma EF_fp: "\<^bold>E\<^bold>F p = p \<union> \<^bold>E\<^bold>X \<^bold>E\<^bold>F p" | |
| 15871 | 128 | proof - | 
| 63055 | 129 | have "mono (\<lambda>s. p \<union> \<^bold>E\<^bold>X s)" by rule auto | 
| 15871 | 130 | then show ?thesis by (simp only: EF_def) (rule lfp_unfold) | 
| 131 | qed | |
| 132 | ||
| 63055 | 133 | lemma AF_fp: "\<^bold>A\<^bold>F p = p \<union> \<^bold>A\<^bold>X \<^bold>A\<^bold>F p" | 
| 15871 | 134 | proof - | 
| 63055 | 135 | have "mono (\<lambda>s. p \<union> \<^bold>A\<^bold>X s)" by rule auto | 
| 15871 | 136 | then show ?thesis by (simp only: AF_lfp) (rule lfp_unfold) | 
| 137 | qed | |
| 138 | ||
| 63055 | 139 | lemma EG_fp: "\<^bold>E\<^bold>G p = p \<inter> \<^bold>E\<^bold>X \<^bold>E\<^bold>G p" | 
| 15871 | 140 | proof - | 
| 63055 | 141 | have "mono (\<lambda>s. p \<inter> \<^bold>E\<^bold>X s)" by rule auto | 
| 15871 | 142 | then show ?thesis by (simp only: EG_def) (rule gfp_unfold) | 
| 143 | qed | |
| 144 | ||
| 61343 | 145 | text \<open> | 
| 69597 | 146 | From the greatest fixed point definition of \<^term>\<open>\<^bold>A\<^bold>G p\<close>, we derive as | 
| 147 | a consequence of the Knaster-Tarski theorem on the one hand that \<^term>\<open>\<^bold>A\<^bold>G p\<close> is a fixed point of the monotonic function | |
| 148 | \<^term>\<open>\<lambda>s. p \<inter> \<^bold>A\<^bold>X s\<close>. | |
| 61343 | 149 | \<close> | 
| 15871 | 150 | |
| 63055 | 151 | lemma AG_fp: "\<^bold>A\<^bold>G p = p \<inter> \<^bold>A\<^bold>X \<^bold>A\<^bold>G p" | 
| 15871 | 152 | proof - | 
| 63055 | 153 | have "mono (\<lambda>s. p \<inter> \<^bold>A\<^bold>X s)" by rule auto | 
| 15871 | 154 | then show ?thesis by (simp only: AG_gfp) (rule gfp_unfold) | 
| 155 | qed | |
| 156 | ||
| 61343 | 157 | text \<open> | 
| 61934 | 158 | This fact may be split up into two inequalities (merely using transitivity | 
| 159 | of \<open>\<subseteq>\<close>, which is an instance of the overloaded \<open>\<le>\<close> in Isabelle/HOL). | |
| 61343 | 160 | \<close> | 
| 15871 | 161 | |
| 63055 | 162 | lemma AG_fp_1: "\<^bold>A\<^bold>G p \<subseteq> p" | 
| 15871 | 163 | proof - | 
| 63055 | 164 | note AG_fp also have "p \<inter> \<^bold>A\<^bold>X \<^bold>A\<^bold>G p \<subseteq> p" by auto | 
| 15871 | 165 | finally show ?thesis . | 
| 166 | qed | |
| 167 | ||
| 63055 | 168 | lemma AG_fp_2: "\<^bold>A\<^bold>G p \<subseteq> \<^bold>A\<^bold>X \<^bold>A\<^bold>G p" | 
| 15871 | 169 | proof - | 
| 63055 | 170 | note AG_fp also have "p \<inter> \<^bold>A\<^bold>X \<^bold>A\<^bold>G p \<subseteq> \<^bold>A\<^bold>X \<^bold>A\<^bold>G p" by auto | 
| 15871 | 171 | finally show ?thesis . | 
| 172 | qed | |
| 173 | ||
| 61343 | 174 | text \<open> | 
| 61934 | 175 | On the other hand, we have from the Knaster-Tarski fixed point theorem that | 
| 69597 | 176 | any other post-fixed point of \<^term>\<open>\<lambda>s. p \<inter> \<^bold>A\<^bold>X s\<close> is smaller than | 
| 177 | \<^term>\<open>\<^bold>A\<^bold>G p\<close>. A post-fixed point is a set of states \<open>q\<close> such that \<^term>\<open>q \<subseteq> p \<inter> \<^bold>A\<^bold>X q\<close>. This leads to the following co-induction principle for | |
| 178 | \<^term>\<open>\<^bold>A\<^bold>G p\<close>. | |
| 61343 | 179 | \<close> | 
| 15871 | 180 | |
| 63055 | 181 | lemma AG_I: "q \<subseteq> p \<inter> \<^bold>A\<^bold>X q \<Longrightarrow> q \<subseteq> \<^bold>A\<^bold>G p" | 
| 15871 | 182 | by (simp only: AG_gfp) (rule gfp_upperbound) | 
| 183 | ||
| 184 | ||
| 61343 | 185 | subsection \<open>The tree induction principle \label{sec:calc-ctl-tree-induct}\<close>
 | 
| 15871 | 186 | |
| 61343 | 187 | text \<open> | 
| 61934 | 188 | With the most basic facts available, we are now able to establish a few more | 
| 63055 | 189 | interesting results, leading to the \<^emph>\<open>tree induction\<close> principle for \<open>\<^bold>A\<^bold>G\<close> | 
| 61934 | 190 | (see below). We will use some elementary monotonicity and distributivity | 
| 191 | rules. | |
| 61343 | 192 | \<close> | 
| 15871 | 193 | |
| 63055 | 194 | lemma AX_int: "\<^bold>A\<^bold>X (p \<inter> q) = \<^bold>A\<^bold>X p \<inter> \<^bold>A\<^bold>X q" by auto | 
| 195 | lemma AX_mono: "p \<subseteq> q \<Longrightarrow> \<^bold>A\<^bold>X p \<subseteq> \<^bold>A\<^bold>X q" by auto | |
| 196 | lemma AG_mono: "p \<subseteq> q \<Longrightarrow> \<^bold>A\<^bold>G p \<subseteq> \<^bold>A\<^bold>G q" | |
| 197 | by (simp only: AG_gfp, rule gfp_mono) auto | |
| 15871 | 198 | |
| 61343 | 199 | text \<open> | 
| 69597 | 200 | The formula \<^term>\<open>AG p\<close> implies \<^term>\<open>AX p\<close> (we use substitution of | 
| 61934 | 201 | \<open>\<subseteq>\<close> with monotonicity). | 
| 61343 | 202 | \<close> | 
| 15871 | 203 | |
| 63055 | 204 | lemma AG_AX: "\<^bold>A\<^bold>G p \<subseteq> \<^bold>A\<^bold>X p" | 
| 15871 | 205 | proof - | 
| 63055 | 206 | have "\<^bold>A\<^bold>G p \<subseteq> \<^bold>A\<^bold>X \<^bold>A\<^bold>G p" by (rule AG_fp_2) | 
| 207 | also have "\<^bold>A\<^bold>G p \<subseteq> p" by (rule AG_fp_1) | |
| 61934 | 208 | moreover note AX_mono | 
| 15871 | 209 | finally show ?thesis . | 
| 210 | qed | |
| 211 | ||
| 61343 | 212 | text \<open> | 
| 63055 | 213 | Furthermore we show idempotency of the \<open>\<^bold>A\<^bold>G\<close> operator. The proof is a good | 
| 61934 | 214 | example of how accumulated facts may get used to feed a single rule step. | 
| 61343 | 215 | \<close> | 
| 15871 | 216 | |
| 63055 | 217 | lemma AG_AG: "\<^bold>A\<^bold>G \<^bold>A\<^bold>G p = \<^bold>A\<^bold>G p" | 
| 15871 | 218 | proof | 
| 63055 | 219 | show "\<^bold>A\<^bold>G \<^bold>A\<^bold>G p \<subseteq> \<^bold>A\<^bold>G p" by (rule AG_fp_1) | 
| 15871 | 220 | next | 
| 63055 | 221 | show "\<^bold>A\<^bold>G p \<subseteq> \<^bold>A\<^bold>G \<^bold>A\<^bold>G p" | 
| 15871 | 222 | proof (rule AG_I) | 
| 63055 | 223 | have "\<^bold>A\<^bold>G p \<subseteq> \<^bold>A\<^bold>G p" .. | 
| 224 | moreover have "\<^bold>A\<^bold>G p \<subseteq> \<^bold>A\<^bold>X \<^bold>A\<^bold>G p" by (rule AG_fp_2) | |
| 225 | ultimately show "\<^bold>A\<^bold>G p \<subseteq> \<^bold>A\<^bold>G p \<inter> \<^bold>A\<^bold>X \<^bold>A\<^bold>G p" .. | |
| 15871 | 226 | qed | 
| 227 | qed | |
| 228 | ||
| 61343 | 229 | text \<open> | 
| 61934 | 230 | \<^smallskip> | 
| 63055 | 231 | We now give an alternative characterization of the \<open>\<^bold>A\<^bold>G\<close> operator, which | 
| 232 | describes the \<open>\<^bold>A\<^bold>G\<close> operator in an ``operational'' way by tree induction: | |
| 69597 | 233 | In a state holds \<^term>\<open>AG p\<close> iff in that state holds \<open>p\<close>, and in all | 
| 63055 | 234 | reachable states \<open>s\<close> follows from the fact that \<open>p\<close> holds in \<open>s\<close>, that \<open>p\<close> | 
| 235 | also holds in all successor states of \<open>s\<close>. We use the co-induction principle | |
| 236 |   @{thm [source] AG_I} to establish this in a purely algebraic manner.
 | |
| 61343 | 237 | \<close> | 
| 15871 | 238 | |
| 63055 | 239 | theorem AG_induct: "p \<inter> \<^bold>A\<^bold>G (p \<rightarrow> \<^bold>A\<^bold>X p) = \<^bold>A\<^bold>G p" | 
| 15871 | 240 | proof | 
| 63055 | 241 | show "p \<inter> \<^bold>A\<^bold>G (p \<rightarrow> \<^bold>A\<^bold>X p) \<subseteq> \<^bold>A\<^bold>G p" (is "?lhs \<subseteq> _") | 
| 15871 | 242 | proof (rule AG_I) | 
| 63055 | 243 | show "?lhs \<subseteq> p \<inter> \<^bold>A\<^bold>X ?lhs" | 
| 15871 | 244 | proof | 
| 245 | show "?lhs \<subseteq> p" .. | |
| 63055 | 246 | show "?lhs \<subseteq> \<^bold>A\<^bold>X ?lhs" | 
| 15871 | 247 | proof - | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32587diff
changeset | 248 |         {
 | 
| 63055 | 249 | have "\<^bold>A\<^bold>G (p \<rightarrow> \<^bold>A\<^bold>X p) \<subseteq> p \<rightarrow> \<^bold>A\<^bold>X p" by (rule AG_fp_1) | 
| 250 | also have "p \<inter> p \<rightarrow> \<^bold>A\<^bold>X p \<subseteq> \<^bold>A\<^bold>X p" .. | |
| 251 | finally have "?lhs \<subseteq> \<^bold>A\<^bold>X p" by auto | |
| 252 | } | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32587diff
changeset | 253 | moreover | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32587diff
changeset | 254 |         {
 | 
| 63055 | 255 | have "p \<inter> \<^bold>A\<^bold>G (p \<rightarrow> \<^bold>A\<^bold>X p) \<subseteq> \<^bold>A\<^bold>G (p \<rightarrow> \<^bold>A\<^bold>X p)" .. | 
| 256 | also have "\<dots> \<subseteq> \<^bold>A\<^bold>X \<dots>" by (rule AG_fp_2) | |
| 257 | finally have "?lhs \<subseteq> \<^bold>A\<^bold>X \<^bold>A\<^bold>G (p \<rightarrow> \<^bold>A\<^bold>X p)" . | |
| 258 | } | |
| 259 | ultimately have "?lhs \<subseteq> \<^bold>A\<^bold>X p \<inter> \<^bold>A\<^bold>X \<^bold>A\<^bold>G (p \<rightarrow> \<^bold>A\<^bold>X p)" .. | |
| 260 | also have "\<dots> = \<^bold>A\<^bold>X ?lhs" by (simp only: AX_int) | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32587diff
changeset | 261 | finally show ?thesis . | 
| 15871 | 262 | qed | 
| 263 | qed | |
| 264 | qed | |
| 265 | next | |
| 63055 | 266 | show "\<^bold>A\<^bold>G p \<subseteq> p \<inter> \<^bold>A\<^bold>G (p \<rightarrow> \<^bold>A\<^bold>X p)" | 
| 15871 | 267 | proof | 
| 63055 | 268 | show "\<^bold>A\<^bold>G p \<subseteq> p" by (rule AG_fp_1) | 
| 269 | show "\<^bold>A\<^bold>G p \<subseteq> \<^bold>A\<^bold>G (p \<rightarrow> \<^bold>A\<^bold>X p)" | |
| 15871 | 270 | proof - | 
| 63055 | 271 | have "\<^bold>A\<^bold>G p = \<^bold>A\<^bold>G \<^bold>A\<^bold>G p" by (simp only: AG_AG) | 
| 272 | also have "\<^bold>A\<^bold>G p \<subseteq> \<^bold>A\<^bold>X p" by (rule AG_AX) moreover note AG_mono | |
| 273 | also have "\<^bold>A\<^bold>X p \<subseteq> (p \<rightarrow> \<^bold>A\<^bold>X p)" .. moreover note AG_mono | |
| 15871 | 274 | finally show ?thesis . | 
| 275 | qed | |
| 276 | qed | |
| 277 | qed | |
| 278 | ||
| 279 | ||
| 61343 | 280 | subsection \<open>An application of tree induction \label{sec:calc-ctl-commute}\<close>
 | 
| 15871 | 281 | |
| 61343 | 282 | text \<open> | 
| 61934 | 283 | Further interesting properties of CTL expressions may be demonstrated with | 
| 63055 | 284 | the help of tree induction; here we show that \<open>\<^bold>A\<^bold>X\<close> and \<open>\<^bold>A\<^bold>G\<close> commute. | 
| 61343 | 285 | \<close> | 
| 15871 | 286 | |
| 63055 | 287 | theorem AG_AX_commute: "\<^bold>A\<^bold>G \<^bold>A\<^bold>X p = \<^bold>A\<^bold>X \<^bold>A\<^bold>G p" | 
| 15871 | 288 | proof - | 
| 63055 | 289 | have "\<^bold>A\<^bold>G \<^bold>A\<^bold>X p = \<^bold>A\<^bold>X p \<inter> \<^bold>A\<^bold>X \<^bold>A\<^bold>G \<^bold>A\<^bold>X p" by (rule AG_fp) | 
| 290 | also have "\<dots> = \<^bold>A\<^bold>X (p \<inter> \<^bold>A\<^bold>G \<^bold>A\<^bold>X p)" by (simp only: AX_int) | |
| 291 | also have "p \<inter> \<^bold>A\<^bold>G \<^bold>A\<^bold>X p = \<^bold>A\<^bold>G p" (is "?lhs = _") | |
| 292 | proof | |
| 293 | have "\<^bold>A\<^bold>X p \<subseteq> p \<rightarrow> \<^bold>A\<^bold>X p" .. | |
| 294 | also have "p \<inter> \<^bold>A\<^bold>G (p \<rightarrow> \<^bold>A\<^bold>X p) = \<^bold>A\<^bold>G p" by (rule AG_induct) | |
| 15871 | 295 | also note Int_mono AG_mono | 
| 63055 | 296 | ultimately show "?lhs \<subseteq> \<^bold>A\<^bold>G p" by fast | 
| 297 | next | |
| 298 | have "\<^bold>A\<^bold>G p \<subseteq> p" by (rule AG_fp_1) | |
| 299 | moreover | |
| 15871 | 300 |     {
 | 
| 63055 | 301 | have "\<^bold>A\<^bold>G p = \<^bold>A\<^bold>G \<^bold>A\<^bold>G p" by (simp only: AG_AG) | 
| 302 | also have "\<^bold>A\<^bold>G p \<subseteq> \<^bold>A\<^bold>X p" by (rule AG_AX) | |
| 15871 | 303 | also note AG_mono | 
| 63055 | 304 | ultimately have "\<^bold>A\<^bold>G p \<subseteq> \<^bold>A\<^bold>G \<^bold>A\<^bold>X p" . | 
| 305 | } | |
| 306 | ultimately show "\<^bold>A\<^bold>G p \<subseteq> ?lhs" .. | |
| 307 | qed | |
| 15871 | 308 | finally show ?thesis . | 
| 309 | qed | |
| 310 | ||
| 311 | end |