doc-src/TutorialI/CTL/CTL.thy
author nipkow
Thu, 14 Sep 2000 17:46:00 +0200
changeset 9958 67f2920862c7
child 9992 4281ccea43f0
permissions -rw-r--r--
*** empty log message ***
nipkow@9958
     1
theory CTL = Main:
nipkow@9958
     2
nipkow@9958
     3
typedecl atom;
nipkow@9958
     4
types state = "atom set";
nipkow@9958
     5
nipkow@9958
     6
datatype ctl_form = Atom atom
nipkow@9958
     7
                  | NOT ctl_form
nipkow@9958
     8
                  | And ctl_form ctl_form
nipkow@9958
     9
                  | AX ctl_form
nipkow@9958
    10
                  | EF ctl_form
nipkow@9958
    11
                  | AF ctl_form;
nipkow@9958
    12
nipkow@9958
    13
consts valid :: "state \<Rightarrow> ctl_form \<Rightarrow> bool" ("(_ \<Turnstile> _)" [80,80] 80)
nipkow@9958
    14
       M :: "(state \<times> state)set";
nipkow@9958
    15
nipkow@9958
    16
constdefs Paths :: "state \<Rightarrow> (nat \<Rightarrow> state)set"
nipkow@9958
    17
"Paths s \<equiv> {p. s = p 0 \<and> (\<forall>i. (p i, p(i+1)) \<in> M)}";
nipkow@9958
    18
nipkow@9958
    19
primrec
nipkow@9958
    20
"s \<Turnstile> Atom a  =  (a\<in>s)"
nipkow@9958
    21
"s \<Turnstile> NOT f   = (~(s \<Turnstile> f))"
nipkow@9958
    22
"s \<Turnstile> And f g = (s \<Turnstile> f \<and> s \<Turnstile> g)"
nipkow@9958
    23
"s \<Turnstile> AX f    = (\<forall>t. (s,t) \<in> M \<longrightarrow> t \<Turnstile> f)"
nipkow@9958
    24
"s \<Turnstile> EF f    = (\<exists>t. (s,t) \<in> M^* \<and> t \<Turnstile> f)"
nipkow@9958
    25
"s \<Turnstile> AF f    = (\<forall>p \<in> Paths s. \<exists>i. p i \<Turnstile> f)";
nipkow@9958
    26
nipkow@9958
    27
constdefs af :: "state set \<Rightarrow> state set \<Rightarrow> state set"
nipkow@9958
    28
"af A T \<equiv> A \<union> {s. \<forall>t. (s, t) \<in> M \<longrightarrow> t \<in> T}";
nipkow@9958
    29
nipkow@9958
    30
lemma mono_af: "mono(af A)";
nipkow@9958
    31
by(force simp add: af_def intro:monoI);
nipkow@9958
    32
nipkow@9958
    33
consts mc :: "ctl_form \<Rightarrow> state set";
nipkow@9958
    34
primrec
nipkow@9958
    35
"mc(Atom a)  = {s. a\<in>s}"
nipkow@9958
    36
"mc(NOT f)   = -mc f"
nipkow@9958
    37
"mc(And f g) = mc f \<inter> mc g"
nipkow@9958
    38
"mc(AX f)    = {s. \<forall>t. (s,t) \<in> M  \<longrightarrow> t \<in> mc f}"
nipkow@9958
    39
"mc(EF f)    = lfp(\<lambda>T. mc f \<union> {s. \<exists>t. (s,t)\<in>M \<and> t\<in>T})"
nipkow@9958
    40
"mc(AF f)    = lfp(af(mc f))";
nipkow@9958
    41
nipkow@9958
    42
lemma mono_ef: "mono(\<lambda>T. A \<union> {s. \<exists>t. (s,t)\<in>M \<and> t\<in>T})";
nipkow@9958
    43
apply(rule monoI);
nipkow@9958
    44
by(blast);
nipkow@9958
    45
nipkow@9958
    46
lemma lfp_conv_EF:
nipkow@9958
    47
"lfp(\<lambda>T. A \<union> {s. \<exists>t. (s,t)\<in>M \<and> t\<in>T}) = {s. \<exists>t. (s,t) \<in> M^* \<and> t \<in> A}";
nipkow@9958
    48
apply(rule equalityI);
nipkow@9958
    49
 apply(rule subsetI);
nipkow@9958
    50
 apply(simp);
nipkow@9958
    51
 apply(erule Lfp.induct);
nipkow@9958
    52
  apply(rule mono_ef);
nipkow@9958
    53
 apply(simp);
nipkow@9958
    54
 apply(blast intro: r_into_rtrancl rtrancl_trans);
nipkow@9958
    55
apply(rule subsetI);
nipkow@9958
    56
apply(simp);
nipkow@9958
    57
apply(erule exE);
nipkow@9958
    58
apply(erule conjE);
nipkow@9958
    59
apply(erule_tac P = "t\<in>A" in rev_mp);
nipkow@9958
    60
apply(erule converse_rtrancl_induct);
nipkow@9958
    61
 apply(rule ssubst [OF lfp_Tarski[OF mono_ef]]);
nipkow@9958
    62
 apply(blast);
nipkow@9958
    63
apply(rule ssubst [OF lfp_Tarski[OF mono_ef]]);
nipkow@9958
    64
by(blast);
nipkow@9958
    65
nipkow@9958
    66
theorem lfp_subset_AF:
nipkow@9958
    67
"lfp(af A) \<subseteq> {s. \<forall> p \<in> Paths s. \<exists> i. p i \<in> A}";
nipkow@9958
    68
apply(rule subsetI);
nipkow@9958
    69
apply(erule Lfp.induct[OF _ mono_af]);
nipkow@9958
    70
apply(simp add: af_def Paths_def);
nipkow@9958
    71
apply(erule disjE);
nipkow@9958
    72
 apply(blast);
nipkow@9958
    73
apply(clarify);
nipkow@9958
    74
apply(erule_tac x = "p 1" in allE);
nipkow@9958
    75
apply(clarsimp);
nipkow@9958
    76
apply(erule_tac x = "\<lambda>i. p(i+1)" in allE);
nipkow@9958
    77
apply(simp);
nipkow@9958
    78
by(blast);
nipkow@9958
    79
nipkow@9958
    80
text{*
nipkow@9958
    81
The opposite direction is proved by contradiction: if some state
nipkow@9958
    82
{term s} is not in @{term"lfp(af A)"}, then we can construct an
nipkow@9958
    83
infinite @{term A}-avoiding path starting from @{term s}. The reason is
nipkow@9958
    84
that by unfolding @{term"lfp"} we find that if @{term s} is not in
nipkow@9958
    85
@{term"lfp(af A)"}, then @{term s} is not in @{term A} and there is a
nipkow@9958
    86
direct successor of @{term s} that is again not in @{term"lfp(af
nipkow@9958
    87
A)"}. Iterating this argument yields the promised infinite
nipkow@9958
    88
@{term A}-avoiding path. Let us formalize this sketch.
nipkow@9958
    89
nipkow@9958
    90
The one-step argument in the above sketch
nipkow@9958
    91
*};
nipkow@9958
    92
nipkow@9958
    93
lemma not_in_lfp_afD:
nipkow@9958
    94
 "s \<notin> lfp(af A) \<Longrightarrow> s \<notin> A \<and> (\<exists> t. (s,t)\<in>M \<and> t \<notin> lfp(af A))";
nipkow@9958
    95
apply(erule swap);
nipkow@9958
    96
apply(rule ssubst[OF lfp_Tarski[OF mono_af]]);
nipkow@9958
    97
by(simp add:af_def);
nipkow@9958
    98
nipkow@9958
    99
text{*\noindent
nipkow@9958
   100
is proved by a variant of contraposition (@{thm[source]swap}:
nipkow@9958
   101
@{thm swap[no_vars]}), i.e.\ assuming the negation of the conclusion
nipkow@9958
   102
and proving @{term"s : lfp(af A)"}. Unfolding @{term lfp} once and
nipkow@9958
   103
simplifying with the definition of @{term af} finishes the proof.
nipkow@9958
   104
nipkow@9958
   105
Now we iterate this process. The following construction of the desired
nipkow@9958
   106
path is parameterized by a predicate @{term P} that should hold along the path:
nipkow@9958
   107
*};
nipkow@9958
   108
nipkow@9958
   109
consts path :: "state \<Rightarrow> (state \<Rightarrow> bool) \<Rightarrow> (nat \<Rightarrow> state)";
nipkow@9958
   110
primrec
nipkow@9958
   111
"path s P 0 = s"
nipkow@9958
   112
"path s P (Suc n) = (SOME t. (path s P n,t) \<in> M \<and> P t)";
nipkow@9958
   113
nipkow@9958
   114
text{*\noindent
nipkow@9958
   115
Element @{term"n+1"} on this path is some arbitrary successor
nipkow@9958
   116
@{term"t"} of element @{term"n"} such that @{term"P t"} holds.  Of
nipkow@9958
   117
course, such a @{term"t"} may in general not exist, but that is of no
nipkow@9958
   118
concern to us since we will only use @{term path} in such cases where a
nipkow@9958
   119
suitable @{term"t"} does exist.
nipkow@9958
   120
nipkow@9958
   121
Now we prove that if each state @{term"s"} that satisfies @{term"P"}
nipkow@9958
   122
has a successor that again satisfies @{term"P"}, then there exists an infinite @{term"P"}-path.
nipkow@9958
   123
*};
nipkow@9958
   124
nipkow@9958
   125
lemma seq_lemma:
nipkow@9958
   126
"\<lbrakk> P s; \<forall>s. P s \<longrightarrow> (\<exists> t. (s,t)\<in>M \<and> P t) \<rbrakk> \<Longrightarrow> \<exists>p\<in>Paths s. \<forall>i. P(p i)";
nipkow@9958
   127
nipkow@9958
   128
txt{*\noindent
nipkow@9958
   129
First we rephrase the conclusion slightly because we need to prove both the path property
nipkow@9958
   130
and the fact that @{term"P"} holds simultaneously:
nipkow@9958
   131
*};
nipkow@9958
   132
nipkow@9958
   133
apply(subgoal_tac "\<exists> p. s = p 0 \<and> (\<forall> i. (p i,p(i+1)) \<in> M \<and> P(p i))");
nipkow@9958
   134
nipkow@9958
   135
txt{*\noindent
nipkow@9958
   136
From this proposition the original goal follows easily
nipkow@9958
   137
*};
nipkow@9958
   138
nipkow@9958
   139
 apply(simp add:Paths_def, blast);
nipkow@9958
   140
apply(rule_tac x = "path s P" in exI);
nipkow@9958
   141
apply(simp);
nipkow@9958
   142
apply(intro strip);
nipkow@9958
   143
apply(induct_tac i);
nipkow@9958
   144
 apply(simp);
nipkow@9958
   145
 apply(fast intro:selectI2EX);
nipkow@9958
   146
apply(simp);
nipkow@9958
   147
apply(rule selectI2EX);
nipkow@9958
   148
 apply(blast);
nipkow@9958
   149
apply(rule selectI2EX);
nipkow@9958
   150
 apply(blast);
nipkow@9958
   151
by(blast);
nipkow@9958
   152
nipkow@9958
   153
lemma seq_lemma:
nipkow@9958
   154
"\<lbrakk> P s; \<forall> s. P s \<longrightarrow> (\<exists> t. (s,t)\<in>M \<and> P t) \<rbrakk> \<Longrightarrow>
nipkow@9958
   155
 \<exists> p\<in>Paths s. \<forall> i. P(p i)";
nipkow@9958
   156
apply(subgoal_tac
nipkow@9958
   157
 "\<exists> p. s = p 0 \<and> (\<forall> i. (p i,p(Suc i))\<in>M \<and> P(p i))");
nipkow@9958
   158
 apply(simp add:Paths_def);
nipkow@9958
   159
 apply(blast);
nipkow@9958
   160
apply(rule_tac x = "nat_rec s (\<lambda>n t. SOME u. (t,u)\<in>M \<and> P u)" in exI);
nipkow@9958
   161
apply(simp);
nipkow@9958
   162
apply(intro strip);
nipkow@9958
   163
apply(induct_tac i);
nipkow@9958
   164
 apply(simp);
nipkow@9958
   165
 apply(fast intro:selectI2EX);
nipkow@9958
   166
apply(simp);
nipkow@9958
   167
apply(rule selectI2EX);
nipkow@9958
   168
 apply(blast);
nipkow@9958
   169
apply(rule selectI2EX);
nipkow@9958
   170
 apply(blast);
nipkow@9958
   171
by(blast);
nipkow@9958
   172
nipkow@9958
   173
theorem AF_subset_lfp:
nipkow@9958
   174
"{s. \<forall> p \<in> Paths s. \<exists> i. p i \<in> A} \<subseteq> lfp(af A)";
nipkow@9958
   175
apply(rule subsetI);
nipkow@9958
   176
apply(erule contrapos2);
nipkow@9958
   177
apply simp;
nipkow@9958
   178
apply(drule seq_lemma);
nipkow@9958
   179
by(auto dest:not_in_lfp_afD);
nipkow@9958
   180
nipkow@9958
   181
nipkow@9958
   182
(*
nipkow@9958
   183
Second proof of opposite direction, directly by wellfounded induction
nipkow@9958
   184
on the initial segment of M that avoids A.
nipkow@9958
   185
nipkow@9958
   186
Avoid s A = the set of successors of s that can be reached by a finite A-avoiding path
nipkow@9958
   187
*)
nipkow@9958
   188
nipkow@9958
   189
consts Avoid :: "state \<Rightarrow> state set \<Rightarrow> state set";
nipkow@9958
   190
inductive "Avoid s A"
nipkow@9958
   191
intros "s \<in> Avoid s A"
nipkow@9958
   192
       "\<lbrakk> t \<in> Avoid s A; t \<notin> A; (t,u) \<in> M \<rbrakk> \<Longrightarrow> u \<in> Avoid s A";
nipkow@9958
   193
nipkow@9958
   194
(* For any infinite A-avoiding path (f) in Avoid s A,
nipkow@9958
   195
   there is some infinite A-avoiding path (p) in Avoid s A that starts with s.
nipkow@9958
   196
*)
nipkow@9958
   197
lemma ex_infinite_path[rule_format]:
nipkow@9958
   198
"t \<in> Avoid s A  \<Longrightarrow>
nipkow@9958
   199
 \<forall>f. t = f 0 \<longrightarrow> (\<forall>i. (f i, f (Suc i)) \<in> M \<and> f i \<in> Avoid s A \<and> f i \<notin> A)
nipkow@9958
   200
                \<longrightarrow> (\<exists> p\<in>Paths s. \<forall>i. p i \<notin> A)";
nipkow@9958
   201
apply(simp add:Paths_def);
nipkow@9958
   202
apply(erule Avoid.induct);
nipkow@9958
   203
 apply(blast);
nipkow@9958
   204
apply(rule allI);
nipkow@9958
   205
apply(erule_tac x = "\<lambda>i. case i of 0 \<Rightarrow> t | Suc i \<Rightarrow> f i" in allE);
nipkow@9958
   206
by(force split:nat.split);
nipkow@9958
   207
nipkow@9958
   208
lemma Avoid_in_lfp[rule_format(no_asm)]:
nipkow@9958
   209
"\<forall>p\<in>Paths s. \<exists>i. p i \<in> A \<Longrightarrow> t \<in> Avoid s A \<longrightarrow> t \<in> lfp(af A)";
nipkow@9958
   210
apply(subgoal_tac "wf{(y,x). (x,y)\<in>M \<and> x \<in> Avoid s A \<and> y \<in> Avoid s A \<and> x \<notin> A}");
nipkow@9958
   211
 apply(erule_tac a = t in wf_induct);
nipkow@9958
   212
 apply(clarsimp);
nipkow@9958
   213
 apply(rule ssubst [OF lfp_Tarski[OF mono_af]]);
nipkow@9958
   214
 apply(unfold af_def);
nipkow@9958
   215
 apply(blast intro:Avoid.intros);
nipkow@9958
   216
apply(erule contrapos2);
nipkow@9958
   217
apply(simp add:wf_iff_no_infinite_down_chain);
nipkow@9958
   218
apply(erule exE);
nipkow@9958
   219
apply(rule ex_infinite_path);
nipkow@9958
   220
by(auto);
nipkow@9958
   221
nipkow@9958
   222
theorem AF_subset_lfp:
nipkow@9958
   223
"{s. \<forall>p \<in> Paths s. \<exists> i. p i \<in> A} \<subseteq> lfp(af A)";
nipkow@9958
   224
apply(rule subsetI);
nipkow@9958
   225
apply(simp);
nipkow@9958
   226
apply(erule Avoid_in_lfp);
nipkow@9958
   227
by(rule Avoid.intros);
nipkow@9958
   228
nipkow@9958
   229
nipkow@9958
   230
theorem "mc f = {s. s \<Turnstile> f}";
nipkow@9958
   231
apply(induct_tac f);
nipkow@9958
   232
by(auto simp add: lfp_conv_EF equalityI[OF lfp_subset_AF AF_subset_lfp]);
nipkow@9958
   233
nipkow@9958
   234
end;