13250
|
1 |
(*<*)theory Sets = Main:(*>*)
|
|
2 |
|
11235
|
3 |
section{*Sets, Functions and Relations*}
|
|
4 |
|
|
5 |
subsection{*Set Notation*}
|
|
6 |
|
13238
|
7 |
text{*
|
|
8 |
\begin{center}
|
|
9 |
\begin{tabular}{ccc}
|
|
10 |
@{term "A \<union> B"} & @{term "A \<inter> B"} & @{term "A - B"} \\
|
|
11 |
@{term "a \<in> A"} & @{term "b \<notin> A"} \\
|
|
12 |
@{term "{a,b}"} & @{text "{x. P x}"} \\
|
|
13 |
@{term "\<Union> M"} & @{text "\<Union>a \<in> A. F a"}
|
|
14 |
\end{tabular}
|
13250
|
15 |
\end{center}*}
|
13249
|
16 |
(*<*)term "A \<union> B" term "A \<inter> B" term "A - B"
|
|
17 |
term "a \<in> A" term "b \<notin> A"
|
|
18 |
term "{a,b}" term "{x. P x}"
|
|
19 |
term "\<Union> M" term "\<Union>a \<in> A. F a"(*>*)
|
13238
|
20 |
|
|
21 |
subsection{*Some Functions*}
|
11235
|
22 |
|
13238
|
23 |
text{*
|
|
24 |
\begin{tabular}{l}
|
|
25 |
@{thm id_def}\\
|
|
26 |
@{thm o_def[no_vars]}\\
|
|
27 |
@{thm image_def[no_vars]}\\
|
|
28 |
@{thm vimage_def[no_vars]}
|
13250
|
29 |
\end{tabular}*}
|
13249
|
30 |
(*<*)thm id_def o_def[no_vars] image_def[no_vars] vimage_def[no_vars](*>*)
|
11235
|
31 |
|
13238
|
32 |
subsection{*Some Relations*}
|
11235
|
33 |
|
13249
|
34 |
text{*
|
|
35 |
\begin{tabular}{l}
|
|
36 |
@{thm Id_def}\\
|
|
37 |
@{thm converse_def[no_vars]}\\
|
|
38 |
@{thm Image_def[no_vars]}\\
|
|
39 |
@{thm rtrancl_refl[no_vars]}\\
|
|
40 |
@{thm rtrancl_into_rtrancl[no_vars]}\\
|
|
41 |
@{thm trancl_def[no_vars]}
|
13250
|
42 |
\end{tabular}*}
|
|
43 |
(*<*)thm Id_def
|
13249
|
44 |
thm converse_def[no_vars]
|
|
45 |
thm Image_def[no_vars]
|
|
46 |
thm relpow.simps[no_vars]
|
|
47 |
thm rtrancl.intros[no_vars]
|
13250
|
48 |
thm trancl_def[no_vars](*>*)
|
11235
|
49 |
|
|
50 |
subsection{*Wellfoundedness*}
|
|
51 |
|
13249
|
52 |
text{*
|
|
53 |
\begin{tabular}{l}
|
|
54 |
@{thm wf_def[no_vars]}\\
|
|
55 |
@{thm wf_iff_no_infinite_down_chain[no_vars]}
|
13250
|
56 |
\end{tabular}*}
|
|
57 |
(*<*)thm wf_def[no_vars]
|
|
58 |
thm wf_iff_no_infinite_down_chain[no_vars](*>*)
|
11235
|
59 |
|
|
60 |
subsection{*Fixed Point Operators*}
|
|
61 |
|
13249
|
62 |
text{*
|
|
63 |
\begin{tabular}{l}
|
|
64 |
@{thm lfp_def[no_vars]}\\
|
|
65 |
@{thm lfp_unfold[no_vars]}\\
|
|
66 |
@{thm lfp_induct[no_vars]}
|
13250
|
67 |
\end{tabular}*}
|
|
68 |
(*<*)thm lfp_def gfp_def
|
11235
|
69 |
thm lfp_unfold
|
13250
|
70 |
thm lfp_induct(*>*)
|
11235
|
71 |
|
|
72 |
subsection{*Case Study: Verified Model Checking*}
|
|
73 |
|
|
74 |
|
|
75 |
typedecl state
|
|
76 |
|
13249
|
77 |
consts M :: "(state \<times> state)set"
|
11235
|
78 |
|
|
79 |
typedecl atom
|
|
80 |
|
|
81 |
consts L :: "state \<Rightarrow> atom set"
|
|
82 |
|
|
83 |
datatype formula = Atom atom
|
|
84 |
| Neg formula
|
|
85 |
| And formula formula
|
|
86 |
| AX formula
|
|
87 |
| EF formula
|
|
88 |
|
|
89 |
consts valid :: "state \<Rightarrow> formula \<Rightarrow> bool" ("(_ \<Turnstile> _)" [80,80] 80)
|
|
90 |
|
|
91 |
primrec
|
|
92 |
"s \<Turnstile> Atom a = (a \<in> L s)"
|
|
93 |
"s \<Turnstile> Neg f = (\<not>(s \<Turnstile> f))"
|
|
94 |
"s \<Turnstile> And f g = (s \<Turnstile> f \<and> s \<Turnstile> g)"
|
|
95 |
"s \<Turnstile> AX f = (\<forall>t. (s,t) \<in> M \<longrightarrow> t \<Turnstile> f)"
|
13249
|
96 |
"s \<Turnstile> EF f = (\<exists>t. (s,t) \<in> M\<^sup>* \<and> t \<Turnstile> f)"
|
11235
|
97 |
|
13249
|
98 |
consts mc :: "formula \<Rightarrow> state set"
|
11235
|
99 |
primrec
|
|
100 |
"mc(Atom a) = {s. a \<in> L s}"
|
|
101 |
"mc(Neg f) = -mc f"
|
|
102 |
"mc(And f g) = mc f \<inter> mc g"
|
|
103 |
"mc(AX f) = {s. \<forall>t. (s,t) \<in> M \<longrightarrow> t \<in> mc f}"
|
|
104 |
"mc(EF f) = lfp(\<lambda>T. mc f \<union> (M\<inverse> `` T))"
|
|
105 |
|
|
106 |
lemma mono_ef: "mono(\<lambda>T. A \<union> (M\<inverse> `` T))"
|
|
107 |
apply(rule monoI)
|
|
108 |
apply blast
|
|
109 |
done
|
|
110 |
|
|
111 |
lemma EF_lemma:
|
|
112 |
"lfp(\<lambda>T. A \<union> (M\<inverse> `` T)) = {s. \<exists>t. (s,t) \<in> M\<^sup>* \<and> t \<in> A}"
|
|
113 |
apply(rule equalityI)
|
|
114 |
thm lfp_lowerbound
|
|
115 |
apply(rule lfp_lowerbound)
|
13249
|
116 |
apply(blast intro: rtrancl_trans)
|
11235
|
117 |
apply(rule subsetI)
|
13249
|
118 |
apply clarsimp
|
11235
|
119 |
apply(erule converse_rtrancl_induct)
|
|
120 |
thm lfp_unfold[OF mono_ef]
|
|
121 |
apply(subst lfp_unfold[OF mono_ef])
|
|
122 |
apply(blast)
|
|
123 |
apply(subst lfp_unfold[OF mono_ef])
|
|
124 |
apply(blast)
|
|
125 |
done
|
|
126 |
|
13249
|
127 |
theorem "mc f = {s. s \<Turnstile> f}"
|
|
128 |
apply(induct_tac f)
|
|
129 |
apply(auto simp add: EF_lemma)
|
|
130 |
done
|
11235
|
131 |
|
|
132 |
text{*
|
|
133 |
\begin{exercise}
|
|
134 |
@{term AX} has a dual operator @{term EN}\footnote{We cannot use the customary @{text EX}
|
|
135 |
as that is the \textsc{ascii}-equivalent of @{text"\<exists>"}}
|
|
136 |
(``there exists a next state such that'') with the intended semantics
|
|
137 |
@{prop[display]"(s \<Turnstile> EN f) = (EX t. (s,t) : M & t \<Turnstile> f)"}
|
|
138 |
Fortunately, @{term"EN f"} can already be expressed as a PDL formula. How?
|
|
139 |
|
|
140 |
Show that the semantics for @{term EF} satisfies the following recursion equation:
|
|
141 |
@{prop[display]"(s \<Turnstile> EF f) = (s \<Turnstile> f | s \<Turnstile> EN(EF f))"}
|
13250
|
142 |
\end{exercise}*}
|
|
143 |
(*<*)end(*>*)
|