diff -r e9ab4ad7bd15 -r 23307fd33906 src/Doc/Tutorial/CTL/CTLind.thy --- a/src/Doc/Tutorial/CTL/CTLind.thy Thu Jan 11 13:48:17 2018 +0100 +++ b/src/Doc/Tutorial/CTL/CTLind.thy Fri Jan 12 14:08:53 2018 +0100 @@ -1,8 +1,8 @@ (*<*)theory CTLind imports CTL begin(*>*) -subsection{*CTL Revisited*} +subsection\CTL Revisited\ -text{*\label{sec:CTL-revisited} +text\\label{sec:CTL-revisited} \index{CTL|(}% The purpose of this section is twofold: to demonstrate some of the induction principles and heuristics discussed above and to @@ -22,7 +22,7 @@ A}-avoiding path: % Second proof of opposite direction, directly by well-founded induction % on the initial segment of M that avoids A. -*} +\ inductive_set Avoid :: "state \ state set \ state set" @@ -31,7 +31,7 @@ "s \ Avoid s A" | "\ t \ Avoid s A; t \ A; (t,u) \ M \ \ u \ Avoid s A" -text{* +text\ It is easy to see that for any infinite @{term A}-avoiding path @{term f} with @{prop"f(0::nat) \ Avoid s A"} there is an infinite @{term A}-avoiding path starting with @{term s} because (by definition of @{const Avoid}) there is a @@ -40,7 +40,7 @@ this requires the following reformulation, as explained in \S\ref{sec:ind-var-in-prems} above; the @{text rule_format} directive undoes the reformulation after the proof. -*} +\ lemma ex_infinite_path[rule_format]: "t \ Avoid s A \ @@ -52,7 +52,7 @@ apply(simp_all add: Paths_def split: nat.split) done -text{*\noindent +text\\noindent The base case (@{prop"t = s"}) is trivial and proved by @{text blast}. In the induction step, we have an infinite @{term A}-avoiding path @{term f} starting from @{term u}, a successor of @{term t}. Now we simply instantiate @@ -66,12 +66,12 @@ inductive proof this must be generalized to the statement that every point @{term t} ``between'' @{term s} and @{term A}, in other words all of @{term"Avoid s A"}, is contained in @{term"lfp(af A)"}: -*} +\ lemma Avoid_in_lfp[rule_format(no_asm)]: "\p\Paths s. \i. p i \ A \ t \ Avoid s A \ t \ lfp(af A)" -txt{*\noindent +txt\\noindent The proof is by induction on the ``distance'' between @{term t} and @{term A}. Remember that @{prop"lfp(af A) = A \ M\ `` lfp(af A)"}. If @{term t} is already in @{term A}, then @{prop"t \ lfp(af A)"} is @@ -85,14 +85,14 @@ As we shall see presently, the absence of infinite @{term A}-avoiding paths starting from @{term s} implies well-foundedness of this relation. For the moment we assume this and proceed with the induction: -*} +\ apply(subgoal_tac "wf{(y,x). (x,y) \ M \ x \ Avoid s A \ x \ A}") apply(erule_tac a = t in wf_induct) apply(clarsimp) (*<*)apply(rename_tac t)(*>*) -txt{*\noindent +txt\\noindent @{subgoals[display,indent=0,margin=65]} Now the induction hypothesis states that if @{prop"t \ A"} then all successors of @{term t} that are in @{term"Avoid s A"} are in @@ -104,13 +104,13 @@ @{term"Avoid s A"}, because we also assume @{prop"t \ Avoid s A"}. Hence, by the induction hypothesis, all successors of @{term t} are indeed in @{term"lfp(af A)"}. Mechanically: -*} +\ apply(subst lfp_unfold[OF mono_af]) apply(simp (no_asm) add: af_def) apply(blast intro: Avoid.intros) -txt{* +txt\ Having proved the main goal, we return to the proof obligation that the relation used above is indeed well-founded. This is proved by contradiction: if the relation is not well-founded then there exists an infinite @{term @@ -119,7 +119,7 @@ @{thm[display]wf_iff_no_infinite_down_chain[no_vars]} From lemma @{thm[source]ex_infinite_path} the existence of an infinite @{term A}-avoiding path starting in @{term s} follows, contradiction. -*} +\ apply(erule contrapos_pp) apply(simp add: wf_iff_no_infinite_down_chain) @@ -128,7 +128,7 @@ apply(auto simp add: Paths_def) done -text{* +text\ The @{text"(no_asm)"} modifier of the @{text"rule_format"} directive in the statement of the lemma means that the assumption is left unchanged; otherwise the @{text"\p"} @@ -139,7 +139,7 @@ The main theorem is simply the corollary where @{prop"t = s"}, when the assumption @{prop"t \ Avoid s A"} is trivially true by the first @{const Avoid}-rule. Isabelle confirms this:% -\index{CTL|)}*} +\index{CTL|)}\ theorem AF_lemma2: "{s. \p \ Paths s. \ i. p i \ A} \ lfp(af A)" by(auto elim: Avoid_in_lfp intro: Avoid.intros)