| 17914 |      1 | (*<*)theory CTLind imports CTL begin(*>*)
 | 
| 10218 |      2 | 
 | 
| 67406 |      3 | subsection\<open>CTL Revisited\<close>
 | 
| 10218 |      4 | 
 | 
| 67406 |      5 | text\<open>\label{sec:CTL-revisited}
 | 
| 11494 |      6 | \index{CTL|(}%
 | 
|  |      7 | The purpose of this section is twofold: to demonstrate
 | 
|  |      8 | some of the induction principles and heuristics discussed above and to
 | 
| 10281 |      9 | show how inductive definitions can simplify proofs.
 | 
| 10218 |     10 | In \S\ref{sec:CTL} we gave a fairly involved proof of the correctness of a
 | 
| 10795 |     11 | model checker for CTL\@. In particular the proof of the
 | 
| 10218 |     12 | @{thm[source]infinity_lemma} on the way to @{thm[source]AF_lemma2} is not as
 | 
| 69505 |     13 | simple as one might expect, due to the \<open>SOME\<close> operator
 | 
| 10281 |     14 | involved. Below we give a simpler proof of @{thm[source]AF_lemma2}
 | 
|  |     15 | based on an auxiliary inductive definition.
 | 
| 10218 |     16 | 
 | 
| 69597 |     17 | Let us call a (finite or infinite) path \emph{\<^term>\<open>A\<close>-avoiding} if it does
 | 
|  |     18 | not touch any node in the set \<^term>\<open>A\<close>. Then @{thm[source]AF_lemma2} says
 | 
|  |     19 | that if no infinite path from some state \<^term>\<open>s\<close> is \<^term>\<open>A\<close>-avoiding,
 | 
|  |     20 | then \<^prop>\<open>s \<in> lfp(af A)\<close>. We prove this by inductively defining the set
 | 
|  |     21 | \<^term>\<open>Avoid s A\<close> of states reachable from \<^term>\<open>s\<close> by a finite \<^term>\<open>A\<close>-avoiding path:
 | 
| 10241 |     22 | % Second proof of opposite direction, directly by well-founded induction
 | 
| 10218 |     23 | % on the initial segment of M that avoids A.
 | 
| 67406 |     24 | \<close>
 | 
| 10218 |     25 | 
 | 
| 23733 |     26 | inductive_set
 | 
|  |     27 |   Avoid :: "state \<Rightarrow> state set \<Rightarrow> state set"
 | 
|  |     28 |   for s :: state and A :: "state set"
 | 
|  |     29 | where
 | 
|  |     30 |     "s \<in> Avoid s A"
 | 
| 58860 |     31 |   | "\<lbrakk> t \<in> Avoid s A; t \<notin> A; (t,u) \<in> M \<rbrakk> \<Longrightarrow> u \<in> Avoid s A"
 | 
| 10218 |     32 | 
 | 
| 67406 |     33 | text\<open>
 | 
| 69597 |     34 | It is easy to see that for any infinite \<^term>\<open>A\<close>-avoiding path \<^term>\<open>f\<close>
 | 
|  |     35 | with \<^prop>\<open>f(0::nat) \<in> Avoid s A\<close> there is an infinite \<^term>\<open>A\<close>-avoiding path
 | 
|  |     36 | starting with \<^term>\<open>s\<close> because (by definition of \<^const>\<open>Avoid\<close>) there is a
 | 
|  |     37 | finite \<^term>\<open>A\<close>-avoiding path from \<^term>\<open>s\<close> to \<^term>\<open>f(0::nat)\<close>.
 | 
|  |     38 | The proof is by induction on \<^prop>\<open>f(0::nat) \<in> Avoid s A\<close>. However,
 | 
| 10218 |     39 | this requires the following
 | 
|  |     40 | reformulation, as explained in \S\ref{sec:ind-var-in-prems} above;
 | 
| 69505 |     41 | the \<open>rule_format\<close> directive undoes the reformulation after the proof.
 | 
| 67406 |     42 | \<close>
 | 
| 10218 |     43 | 
 | 
|  |     44 | lemma ex_infinite_path[rule_format]:
 | 
|  |     45 |   "t \<in> Avoid s A  \<Longrightarrow>
 | 
| 58860 |     46 |    \<forall>f\<in>Paths t. (\<forall>i. f i \<notin> A) \<longrightarrow> (\<exists>p\<in>Paths s. \<forall>i. p i \<notin> A)"
 | 
|  |     47 | apply(erule Avoid.induct)
 | 
|  |     48 |  apply(blast)
 | 
|  |     49 | apply(clarify)
 | 
|  |     50 | apply(drule_tac x = "\<lambda>i. case i of 0 \<Rightarrow> t | Suc i \<Rightarrow> f i" in bspec)
 | 
|  |     51 | apply(simp_all add: Paths_def split: nat.split)
 | 
| 10218 |     52 | done
 | 
|  |     53 | 
 | 
| 67406 |     54 | text\<open>\noindent
 | 
| 69597 |     55 | The base case (\<^prop>\<open>t = s\<close>) is trivial and proved by \<open>blast\<close>.
 | 
|  |     56 | In the induction step, we have an infinite \<^term>\<open>A\<close>-avoiding path \<^term>\<open>f\<close>
 | 
|  |     57 | starting from \<^term>\<open>u\<close>, a successor of \<^term>\<open>t\<close>. Now we simply instantiate
 | 
| 69505 |     58 | the \<open>\<forall>f\<in>Paths t\<close> in the induction hypothesis by the path starting with
 | 
| 69597 |     59 | \<^term>\<open>t\<close> and continuing with \<^term>\<open>f\<close>. That is what the above $\lambda$-term
 | 
|  |     60 | expresses.  Simplification shows that this is a path starting with \<^term>\<open>t\<close> 
 | 
| 10885 |     61 | and that the instantiated induction hypothesis implies the conclusion.
 | 
| 10218 |     62 | 
 | 
| 69597 |     63 | Now we come to the key lemma. Assuming that no infinite \<^term>\<open>A\<close>-avoiding
 | 
|  |     64 | path starts from \<^term>\<open>s\<close>, we want to show \<^prop>\<open>s \<in> lfp(af A)\<close>. For the
 | 
|  |     65 | inductive proof this must be generalized to the statement that every point \<^term>\<open>t\<close>
 | 
|  |     66 | ``between'' \<^term>\<open>s\<close> and \<^term>\<open>A\<close>, in other words all of \<^term>\<open>Avoid s A\<close>,
 | 
|  |     67 | is contained in \<^term>\<open>lfp(af A)\<close>:
 | 
| 67406 |     68 | \<close>
 | 
| 10218 |     69 | 
 | 
|  |     70 | lemma Avoid_in_lfp[rule_format(no_asm)]:
 | 
| 58860 |     71 |   "\<forall>p\<in>Paths s. \<exists>i. p i \<in> A \<Longrightarrow> t \<in> Avoid s A \<longrightarrow> t \<in> lfp(af A)"
 | 
| 11196 |     72 | 
 | 
| 67406 |     73 | txt\<open>\noindent
 | 
| 69597 |     74 | The proof is by induction on the ``distance'' between \<^term>\<open>t\<close> and \<^term>\<open>A\<close>. Remember that \<^prop>\<open>lfp(af A) = A \<union> M\<inverse> `` lfp(af A)\<close>.
 | 
|  |     75 | If \<^term>\<open>t\<close> is already in \<^term>\<open>A\<close>, then \<^prop>\<open>t \<in> lfp(af A)\<close> is
 | 
|  |     76 | trivial. If \<^term>\<open>t\<close> is not in \<^term>\<open>A\<close> but all successors are in
 | 
|  |     77 | \<^term>\<open>lfp(af A)\<close> (induction hypothesis), then \<^prop>\<open>t \<in> lfp(af A)\<close> is
 | 
| 11196 |     78 | again trivial.
 | 
|  |     79 | 
 | 
|  |     80 | The formal counterpart of this proof sketch is a well-founded induction
 | 
| 69597 |     81 | on~\<^term>\<open>M\<close> restricted to \<^term>\<open>Avoid s A - A\<close>, roughly speaking:
 | 
| 11196 |     82 | @{term[display]"{(y,x). (x,y) \<in> M \<and> x \<in> Avoid s A \<and> x \<notin> A}"}
 | 
| 69597 |     83 | As we shall see presently, the absence of infinite \<^term>\<open>A\<close>-avoiding paths
 | 
|  |     84 | starting from \<^term>\<open>s\<close> implies well-foundedness of this relation. For the
 | 
| 10218 |     85 | moment we assume this and proceed with the induction:
 | 
| 67406 |     86 | \<close>
 | 
| 10218 |     87 | 
 | 
| 58860 |     88 | apply(subgoal_tac "wf{(y,x). (x,y) \<in> M \<and> x \<in> Avoid s A \<and> x \<notin> A}")
 | 
|  |     89 |  apply(erule_tac a = t in wf_induct)
 | 
|  |     90 |  apply(clarsimp)
 | 
| 11196 |     91 | (*<*)apply(rename_tac t)(*>*)
 | 
| 10218 |     92 | 
 | 
| 67406 |     93 | txt\<open>\noindent
 | 
| 10885 |     94 | @{subgoals[display,indent=0,margin=65]}
 | 
| 69597 |     95 | Now the induction hypothesis states that if \<^prop>\<open>t \<notin> A\<close>
 | 
|  |     96 | then all successors of \<^term>\<open>t\<close> that are in \<^term>\<open>Avoid s A\<close> are in
 | 
|  |     97 | \<^term>\<open>lfp (af A)\<close>. Unfolding \<^term>\<open>lfp\<close> in the conclusion of the first
 | 
|  |     98 | subgoal once, we have to prove that \<^term>\<open>t\<close> is in \<^term>\<open>A\<close> or all successors
 | 
|  |     99 | of \<^term>\<open>t\<close> are in \<^term>\<open>lfp (af A)\<close>.  But if \<^term>\<open>t\<close> is not in \<^term>\<open>A\<close>,
 | 
| 11196 |    100 | the second 
 | 
| 69597 |    101 | \<^const>\<open>Avoid\<close>-rule implies that all successors of \<^term>\<open>t\<close> are in
 | 
|  |    102 | \<^term>\<open>Avoid s A\<close>, because we also assume \<^prop>\<open>t \<in> Avoid s A\<close>.
 | 
|  |    103 | Hence, by the induction hypothesis, all successors of \<^term>\<open>t\<close> are indeed in
 | 
|  |    104 | \<^term>\<open>lfp(af A)\<close>. Mechanically:
 | 
| 67406 |    105 | \<close>
 | 
| 10218 |    106 | 
 | 
| 58860 |    107 |  apply(subst lfp_unfold[OF mono_af])
 | 
|  |    108 |  apply(simp (no_asm) add: af_def)
 | 
|  |    109 |  apply(blast intro: Avoid.intros)
 | 
| 10218 |    110 | 
 | 
| 67406 |    111 | txt\<open>
 | 
| 11494 |    112 | Having proved the main goal, we return to the proof obligation that the 
 | 
|  |    113 | relation used above is indeed well-founded. This is proved by contradiction: if
 | 
| 69597 |    114 | the relation is not well-founded then there exists an infinite \<^term>\<open>A\<close>-avoiding path all in \<^term>\<open>Avoid s A\<close>, by theorem
 | 
| 10218 |    115 | @{thm[source]wf_iff_no_infinite_down_chain}:
 | 
|  |    116 | @{thm[display]wf_iff_no_infinite_down_chain[no_vars]}
 | 
|  |    117 | From lemma @{thm[source]ex_infinite_path} the existence of an infinite
 | 
| 69597 |    118 | \<^term>\<open>A\<close>-avoiding path starting in \<^term>\<open>s\<close> follows, contradiction.
 | 
| 67406 |    119 | \<close>
 | 
| 10218 |    120 | 
 | 
| 58860 |    121 | apply(erule contrapos_pp)
 | 
|  |    122 | apply(simp add: wf_iff_no_infinite_down_chain)
 | 
|  |    123 | apply(erule exE)
 | 
|  |    124 | apply(rule ex_infinite_path)
 | 
|  |    125 | apply(auto simp add: Paths_def)
 | 
| 10218 |    126 | done
 | 
|  |    127 | 
 | 
| 67406 |    128 | text\<open>
 | 
| 69505 |    129 | The \<open>(no_asm)\<close> modifier of the \<open>rule_format\<close> directive in the
 | 
| 11196 |    130 | statement of the lemma means
 | 
| 69505 |    131 | that the assumption is left unchanged; otherwise the \<open>\<forall>p\<close> 
 | 
| 11494 |    132 | would be turned
 | 
| 69505 |    133 | into a \<open>\<And>p\<close>, which would complicate matters below. As it is,
 | 
| 10218 |    134 | @{thm[source]Avoid_in_lfp} is now
 | 
|  |    135 | @{thm[display]Avoid_in_lfp[no_vars]}
 | 
| 69597 |    136 | The main theorem is simply the corollary where \<^prop>\<open>t = s\<close>,
 | 
|  |    137 | when the assumption \<^prop>\<open>t \<in> Avoid s A\<close> is trivially true
 | 
|  |    138 | by the first \<^const>\<open>Avoid\<close>-rule. Isabelle confirms this:%
 | 
| 67406 |    139 | \index{CTL|)}\<close>
 | 
| 10218 |    140 | 
 | 
| 58860 |    141 | theorem AF_lemma2:  "{s. \<forall>p \<in> Paths s. \<exists> i. p i \<in> A} \<subseteq> lfp(af A)"
 | 
|  |    142 | by(auto elim: Avoid_in_lfp intro: Avoid.intros)
 | 
| 10218 |    143 | 
 | 
|  |    144 | 
 | 
|  |    145 | (*<*)end(*>*)
 |