8745
|
1 |
(*<*)
|
16417
|
2 |
theory Ifexpr imports Main begin;
|
8745
|
3 |
(*>*)
|
|
4 |
|
10885
|
5 |
subsection{*Case Study: Boolean Expressions*}
|
9933
|
6 |
|
11456
|
7 |
text{*\label{sec:boolex}\index{boolean expressions example|(}
|
9933
|
8 |
The aim of this case study is twofold: it shows how to model boolean
|
|
9 |
expressions and some algorithms for manipulating them, and it demonstrates
|
|
10 |
the constructs introduced above.
|
|
11 |
*}
|
|
12 |
|
10978
|
13 |
subsubsection{*Modelling Boolean Expressions*}
|
9933
|
14 |
|
8745
|
15 |
text{*
|
|
16 |
We want to represent boolean expressions built up from variables and
|
|
17 |
constants by negation and conjunction. The following datatype serves exactly
|
|
18 |
that purpose:
|
|
19 |
*}
|
|
20 |
|
|
21 |
datatype boolex = Const bool | Var nat | Neg boolex
|
|
22 |
| And boolex boolex;
|
|
23 |
|
|
24 |
text{*\noindent
|
9541
|
25 |
The two constants are represented by @{term"Const True"} and
|
|
26 |
@{term"Const False"}. Variables are represented by terms of the form
|
9792
|
27 |
@{term"Var n"}, where @{term"n"} is a natural number (type @{typ"nat"}).
|
8745
|
28 |
For example, the formula $P@0 \land \neg P@1$ is represented by the term
|
9541
|
29 |
@{term"And (Var 0) (Neg(Var 1))"}.
|
8745
|
30 |
|
10978
|
31 |
\subsubsection{The Value of a Boolean Expression}
|
8745
|
32 |
|
|
33 |
The value of a boolean expression depends on the value of its variables.
|
9792
|
34 |
Hence the function @{text"value"} takes an additional parameter, an
|
|
35 |
\emph{environment} of type @{typ"nat => bool"}, which maps variables to their
|
|
36 |
values:
|
8745
|
37 |
*}
|
|
38 |
|
17555
|
39 |
consts "value" :: "boolex \<Rightarrow> (nat \<Rightarrow> bool) \<Rightarrow> bool";
|
8745
|
40 |
primrec
|
|
41 |
"value (Const b) env = b"
|
|
42 |
"value (Var x) env = env x"
|
10795
|
43 |
"value (Neg b) env = (\<not> value b env)"
|
|
44 |
"value (And b c) env = (value b env \<and> value c env)";
|
8745
|
45 |
|
|
46 |
text{*\noindent
|
10885
|
47 |
\subsubsection{If-Expressions}
|
8745
|
48 |
|
|
49 |
An alternative and often more efficient (because in a certain sense
|
|
50 |
canonical) representation are so-called \emph{If-expressions} built up
|
9792
|
51 |
from constants (@{term"CIF"}), variables (@{term"VIF"}) and conditionals
|
|
52 |
(@{term"IF"}):
|
8745
|
53 |
*}
|
|
54 |
|
|
55 |
datatype ifex = CIF bool | VIF nat | IF ifex ifex ifex;
|
|
56 |
|
|
57 |
text{*\noindent
|
10971
|
58 |
The evaluation of If-expressions proceeds as for @{typ"boolex"}:
|
8745
|
59 |
*}
|
|
60 |
|
10795
|
61 |
consts valif :: "ifex \<Rightarrow> (nat \<Rightarrow> bool) \<Rightarrow> bool";
|
8745
|
62 |
primrec
|
|
63 |
"valif (CIF b) env = b"
|
|
64 |
"valif (VIF x) env = env x"
|
|
65 |
"valif (IF b t e) env = (if valif b env then valif t env
|
|
66 |
else valif e env)";
|
|
67 |
|
|
68 |
text{*
|
10978
|
69 |
\subsubsection{Converting Boolean and If-Expressions}
|
8745
|
70 |
|
9792
|
71 |
The type @{typ"boolex"} is close to the customary representation of logical
|
|
72 |
formulae, whereas @{typ"ifex"} is designed for efficiency. It is easy to
|
|
73 |
translate from @{typ"boolex"} into @{typ"ifex"}:
|
8745
|
74 |
*}
|
|
75 |
|
10795
|
76 |
consts bool2if :: "boolex \<Rightarrow> ifex";
|
8745
|
77 |
primrec
|
|
78 |
"bool2if (Const b) = CIF b"
|
|
79 |
"bool2if (Var x) = VIF x"
|
|
80 |
"bool2if (Neg b) = IF (bool2if b) (CIF False) (CIF True)"
|
|
81 |
"bool2if (And b c) = IF (bool2if b) (bool2if c) (CIF False)";
|
|
82 |
|
|
83 |
text{*\noindent
|
9792
|
84 |
At last, we have something we can verify: that @{term"bool2if"} preserves the
|
8745
|
85 |
value of its argument:
|
|
86 |
*}
|
|
87 |
|
|
88 |
lemma "valif (bool2if b) env = value b env";
|
|
89 |
|
|
90 |
txt{*\noindent
|
|
91 |
The proof is canonical:
|
|
92 |
*}
|
|
93 |
|
|
94 |
apply(induct_tac b);
|
10171
|
95 |
apply(auto);
|
|
96 |
done
|
8745
|
97 |
|
|
98 |
text{*\noindent
|
|
99 |
In fact, all proofs in this case study look exactly like this. Hence we do
|
|
100 |
not show them below.
|
|
101 |
|
|
102 |
More interesting is the transformation of If-expressions into a normal form
|
9792
|
103 |
where the first argument of @{term"IF"} cannot be another @{term"IF"} but
|
8745
|
104 |
must be a constant or variable. Such a normal form can be computed by
|
9541
|
105 |
repeatedly replacing a subterm of the form @{term"IF (IF b x y) z u"} by
|
|
106 |
@{term"IF b (IF x z u) (IF y z u)"}, which has the same value. The following
|
8745
|
107 |
primitive recursive functions perform this task:
|
|
108 |
*}
|
|
109 |
|
10795
|
110 |
consts normif :: "ifex \<Rightarrow> ifex \<Rightarrow> ifex \<Rightarrow> ifex";
|
8745
|
111 |
primrec
|
|
112 |
"normif (CIF b) t e = IF (CIF b) t e"
|
|
113 |
"normif (VIF x) t e = IF (VIF x) t e"
|
|
114 |
"normif (IF b t e) u f = normif b (normif t u f) (normif e u f)";
|
|
115 |
|
10795
|
116 |
consts norm :: "ifex \<Rightarrow> ifex";
|
8745
|
117 |
primrec
|
|
118 |
"norm (CIF b) = CIF b"
|
|
119 |
"norm (VIF x) = VIF x"
|
|
120 |
"norm (IF b t e) = normif b (norm t) (norm e)";
|
|
121 |
|
|
122 |
text{*\noindent
|
11456
|
123 |
Their interplay is tricky; we leave it to you to develop an
|
8745
|
124 |
intuitive understanding. Fortunately, Isabelle can help us to verify that the
|
|
125 |
transformation preserves the value of the expression:
|
|
126 |
*}
|
|
127 |
|
|
128 |
theorem "valif (norm b) env = valif b env";(*<*)oops;(*>*)
|
|
129 |
|
|
130 |
text{*\noindent
|
|
131 |
The proof is canonical, provided we first show the following simplification
|
11456
|
132 |
lemma, which also helps to understand what @{term"normif"} does:
|
8745
|
133 |
*}
|
|
134 |
|
|
135 |
lemma [simp]:
|
10795
|
136 |
"\<forall>t e. valif (normif b t e) env = valif (IF b t e) env";
|
8745
|
137 |
(*<*)
|
|
138 |
apply(induct_tac b);
|
9458
|
139 |
by(auto);
|
8745
|
140 |
|
|
141 |
theorem "valif (norm b) env = valif b env";
|
|
142 |
apply(induct_tac b);
|
9458
|
143 |
by(auto);
|
8745
|
144 |
(*>*)
|
|
145 |
text{*\noindent
|
|
146 |
Note that the lemma does not have a name, but is implicitly used in the proof
|
9792
|
147 |
of the theorem shown above because of the @{text"[simp]"} attribute.
|
8745
|
148 |
|
9792
|
149 |
But how can we be sure that @{term"norm"} really produces a normal form in
|
11456
|
150 |
the above sense? We define a function that tests If-expressions for normality:
|
8745
|
151 |
*}
|
|
152 |
|
10795
|
153 |
consts normal :: "ifex \<Rightarrow> bool";
|
8745
|
154 |
primrec
|
|
155 |
"normal(CIF b) = True"
|
|
156 |
"normal(VIF x) = True"
|
10795
|
157 |
"normal(IF b t e) = (normal t \<and> normal e \<and>
|
|
158 |
(case b of CIF b \<Rightarrow> True | VIF x \<Rightarrow> True | IF x y z \<Rightarrow> False))";
|
8745
|
159 |
|
|
160 |
text{*\noindent
|
11456
|
161 |
Now we prove @{term"normal(norm b)"}. Of course, this requires a lemma about
|
9792
|
162 |
normality of @{term"normif"}:
|
8745
|
163 |
*}
|
|
164 |
|
12327
|
165 |
lemma [simp]: "\<forall>t e. normal(normif b t e) = (normal t \<and> normal e)";
|
8771
|
166 |
(*<*)
|
8745
|
167 |
apply(induct_tac b);
|
9458
|
168 |
by(auto);
|
8745
|
169 |
|
|
170 |
theorem "normal(norm b)";
|
|
171 |
apply(induct_tac b);
|
9458
|
172 |
by(auto);
|
9933
|
173 |
(*>*)
|
8745
|
174 |
|
9933
|
175 |
text{*\medskip
|
10795
|
176 |
How do we come up with the required lemmas? Try to prove the main theorems
|
|
177 |
without them and study carefully what @{text auto} leaves unproved. This
|
|
178 |
can provide the clue. The necessity of universal quantification
|
9933
|
179 |
(@{text"\<forall>t e"}) in the two lemmas is explained in
|
|
180 |
\S\ref{sec:InductionHeuristics}
|
|
181 |
|
|
182 |
\begin{exercise}
|
15905
|
183 |
We strengthen the definition of a @{const normal} If-expression as follows:
|
9933
|
184 |
the first argument of all @{term IF}s must be a variable. Adapt the above
|
|
185 |
development to this changed requirement. (Hint: you may need to formulate
|
|
186 |
some of the goals as implications (@{text"\<longrightarrow>"}) rather than
|
|
187 |
equalities (@{text"="}).)
|
|
188 |
\end{exercise}
|
11456
|
189 |
\index{boolean expressions example|)}
|
9933
|
190 |
*}
|
|
191 |
(*<*)
|
15243
|
192 |
|
|
193 |
consts normif2 :: "ifex => ifex => ifex => ifex"
|
|
194 |
primrec
|
|
195 |
"normif2 (CIF b) t e = (if b then t else e)"
|
|
196 |
"normif2 (VIF x) t e = IF (VIF x) t e"
|
|
197 |
"normif2 (IF b t e) u f = normif2 b (normif2 t u f) (normif2 e u f)"
|
|
198 |
|
|
199 |
consts norm2 :: "ifex => ifex"
|
|
200 |
primrec
|
|
201 |
"norm2 (CIF b) = CIF b"
|
|
202 |
"norm2 (VIF x) = VIF x"
|
|
203 |
"norm2 (IF b t e) = normif2 b (norm2 t) (norm2 e)"
|
|
204 |
|
|
205 |
consts normal2 :: "ifex => bool"
|
|
206 |
primrec
|
|
207 |
"normal2(CIF b) = True"
|
|
208 |
"normal2(VIF x) = True"
|
|
209 |
"normal2(IF b t e) = (normal2 t & normal2 e &
|
|
210 |
(case b of CIF b => False | VIF x => True | IF x y z => False))"
|
|
211 |
|
|
212 |
|
|
213 |
lemma [simp]:
|
|
214 |
"ALL t e. valif (normif2 b t e) env = valif (IF b t e) env"
|
|
215 |
apply(induct b)
|
|
216 |
by(auto)
|
|
217 |
|
|
218 |
theorem "valif (norm2 b) env = valif b env"
|
|
219 |
apply(induct b)
|
|
220 |
by(auto)
|
|
221 |
|
|
222 |
lemma [simp]: "ALL t e. normal2 t & normal2 e --> normal2(normif2 b t e)"
|
|
223 |
apply(induct b)
|
|
224 |
by(auto)
|
|
225 |
|
|
226 |
theorem "normal2(norm2 b)"
|
|
227 |
apply(induct b)
|
|
228 |
by(auto)
|
|
229 |
|
8771
|
230 |
end
|
|
231 |
(*>*)
|