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