doc-src/TutorialI/Ifexpr/Ifexpr.thy
author chaieb
Wed, 19 May 2004 11:23:59 +0200
changeset 14758 af3b71a46a1c
parent 12327 5a4d78204492
child 15243 ba52fdc2c4e8
permissions -rw-r--r--
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller. the tactic has also changed and allows the abstaction over fuction occurences whose type is nat or int.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     1
(*<*)
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     2
theory Ifexpr = Main:;
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     3
(*>*)
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     4
10885
90695f46440b lcp's pass over the book, chapters 1-8
paulson
parents: 10795
diff changeset
     5
subsection{*Case Study: Boolean Expressions*}
9933
9feb1e0c4cb3 *** empty log message ***
nipkow
parents: 9792
diff changeset
     6
11456
7eb63f63e6c6 revisions and indexing
paulson
parents: 10978
diff changeset
     7
text{*\label{sec:boolex}\index{boolean expressions example|(}
9933
9feb1e0c4cb3 *** empty log message ***
nipkow
parents: 9792
diff changeset
     8
The aim of this case study is twofold: it shows how to model boolean
9feb1e0c4cb3 *** empty log message ***
nipkow
parents: 9792
diff changeset
     9
expressions and some algorithms for manipulating them, and it demonstrates
9feb1e0c4cb3 *** empty log message ***
nipkow
parents: 9792
diff changeset
    10
the constructs introduced above.
9feb1e0c4cb3 *** empty log message ***
nipkow
parents: 9792
diff changeset
    11
*}
9feb1e0c4cb3 *** empty log message ***
nipkow
parents: 9792
diff changeset
    12
10978
5eebea8f359f *** empty log message ***
nipkow
parents: 10971
diff changeset
    13
subsubsection{*Modelling Boolean Expressions*}
9933
9feb1e0c4cb3 *** empty log message ***
nipkow
parents: 9792
diff changeset
    14
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    15
text{*
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    16
We want to represent boolean expressions built up from variables and
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    17
constants by negation and conjunction. The following datatype serves exactly
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    18
that purpose:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    19
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    20
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    21
datatype boolex = Const bool | Var nat | Neg boolex
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    22
                | And boolex boolex;
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    23
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    24
text{*\noindent
9541
d17c0b34d5c8 *** empty log message ***
nipkow
parents: 9458
diff changeset
    25
The two constants are represented by @{term"Const True"} and
d17c0b34d5c8 *** empty log message ***
nipkow
parents: 9458
diff changeset
    26
@{term"Const False"}. Variables are represented by terms of the form
9792
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9541
diff changeset
    27
@{term"Var n"}, where @{term"n"} is a natural number (type @{typ"nat"}).
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    28
For example, the formula $P@0 \land \neg P@1$ is represented by the term
9541
d17c0b34d5c8 *** empty log message ***
nipkow
parents: 9458
diff changeset
    29
@{term"And (Var 0) (Neg(Var 1))"}.
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    30
10978
5eebea8f359f *** empty log message ***
nipkow
parents: 10971
diff changeset
    31
\subsubsection{The Value of a Boolean Expression}
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    32
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    33
The value of a boolean expression depends on the value of its variables.
9792
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9541
diff changeset
    34
Hence the function @{text"value"} takes an additional parameter, an
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9541
diff changeset
    35
\emph{environment} of type @{typ"nat => bool"}, which maps variables to their
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9541
diff changeset
    36
values:
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    37
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    38
10795
9e888d60d3e5 minor edits to Chapters 1-3
paulson
parents: 10171
diff changeset
    39
consts value :: "boolex \<Rightarrow> (nat \<Rightarrow> bool) \<Rightarrow> bool";
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    40
primrec
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    41
"value (Const b) env = b"
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    42
"value (Var x)   env = env x"
10795
9e888d60d3e5 minor edits to Chapters 1-3
paulson
parents: 10171
diff changeset
    43
"value (Neg b)   env = (\<not> value b env)"
9e888d60d3e5 minor edits to Chapters 1-3
paulson
parents: 10171
diff changeset
    44
"value (And b c) env = (value b env \<and> value c env)";
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    45
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    46
text{*\noindent
10885
90695f46440b lcp's pass over the book, chapters 1-8
paulson
parents: 10795
diff changeset
    47
\subsubsection{If-Expressions}
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    48
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    49
An alternative and often more efficient (because in a certain sense
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    50
canonical) representation are so-called \emph{If-expressions} built up
9792
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9541
diff changeset
    51
from constants (@{term"CIF"}), variables (@{term"VIF"}) and conditionals
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9541
diff changeset
    52
(@{term"IF"}):
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    53
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    54
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    55
datatype ifex = CIF bool | VIF nat | IF ifex ifex ifex;
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    56
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    57
text{*\noindent
10971
6852682eaf16 *** empty log message ***
nipkow
parents: 10885
diff changeset
    58
The evaluation of If-expressions proceeds as for @{typ"boolex"}:
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    59
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    60
10795
9e888d60d3e5 minor edits to Chapters 1-3
paulson
parents: 10171
diff changeset
    61
consts valif :: "ifex \<Rightarrow> (nat \<Rightarrow> bool) \<Rightarrow> bool";
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    62
primrec
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    63
"valif (CIF b)    env = b"
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    64
"valif (VIF x)    env = env x"
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    65
"valif (IF b t e) env = (if valif b env then valif t env
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    66
                                        else valif e env)";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    67
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    68
text{*
10978
5eebea8f359f *** empty log message ***
nipkow
parents: 10971
diff changeset
    69
\subsubsection{Converting Boolean and If-Expressions}
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    70
9792
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9541
diff changeset
    71
The type @{typ"boolex"} is close to the customary representation of logical
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9541
diff changeset
    72
formulae, whereas @{typ"ifex"} is designed for efficiency. It is easy to
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9541
diff changeset
    73
translate from @{typ"boolex"} into @{typ"ifex"}:
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    74
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    75
10795
9e888d60d3e5 minor edits to Chapters 1-3
paulson
parents: 10171
diff changeset
    76
consts bool2if :: "boolex \<Rightarrow> ifex";
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    77
primrec
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    78
"bool2if (Const b) = CIF b"
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    79
"bool2if (Var x)   = VIF x"
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    80
"bool2if (Neg b)   = IF (bool2if b) (CIF False) (CIF True)"
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    81
"bool2if (And b c) = IF (bool2if b) (bool2if c) (CIF False)";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    82
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    83
text{*\noindent
9792
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9541
diff changeset
    84
At last, we have something we can verify: that @{term"bool2if"} preserves the
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    85
value of its argument:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    86
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    87
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    88
lemma "valif (bool2if b) env = value b env";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    89
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    90
txt{*\noindent
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    91
The proof is canonical:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    92
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    93
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    94
apply(induct_tac b);
10171
59d6633835fa *** empty log message ***
nipkow
parents: 9933
diff changeset
    95
apply(auto);
59d6633835fa *** empty log message ***
nipkow
parents: 9933
diff changeset
    96
done
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    97
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    98
text{*\noindent
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    99
In fact, all proofs in this case study look exactly like this. Hence we do
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   100
not show them below.
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   101
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   102
More interesting is the transformation of If-expressions into a normal form
9792
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9541
diff changeset
   103
where the first argument of @{term"IF"} cannot be another @{term"IF"} but
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   104
must be a constant or variable. Such a normal form can be computed by
9541
d17c0b34d5c8 *** empty log message ***
nipkow
parents: 9458
diff changeset
   105
repeatedly replacing a subterm of the form @{term"IF (IF b x y) z u"} by
d17c0b34d5c8 *** empty log message ***
nipkow
parents: 9458
diff changeset
   106
@{term"IF b (IF x z u) (IF y z u)"}, which has the same value. The following
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   107
primitive recursive functions perform this task:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   108
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   109
10795
9e888d60d3e5 minor edits to Chapters 1-3
paulson
parents: 10171
diff changeset
   110
consts normif :: "ifex \<Rightarrow> ifex \<Rightarrow> ifex \<Rightarrow> ifex";
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   111
primrec
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   112
"normif (CIF b)    t e = IF (CIF b) t e"
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   113
"normif (VIF x)    t e = IF (VIF x) t e"
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   114
"normif (IF b t e) u f = normif b (normif t u f) (normif e u f)";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   115
10795
9e888d60d3e5 minor edits to Chapters 1-3
paulson
parents: 10171
diff changeset
   116
consts norm :: "ifex \<Rightarrow> ifex";
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   117
primrec
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   118
"norm (CIF b)    = CIF b"
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   119
"norm (VIF x)    = VIF x"
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   120
"norm (IF b t e) = normif b (norm t) (norm e)";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   121
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   122
text{*\noindent
11456
7eb63f63e6c6 revisions and indexing
paulson
parents: 10978
diff changeset
   123
Their interplay is tricky; we leave it to you to develop an
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   124
intuitive understanding. Fortunately, Isabelle can help us to verify that the
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   125
transformation preserves the value of the expression:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   126
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   127
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   128
theorem "valif (norm b) env = valif b env";(*<*)oops;(*>*)
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   129
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   130
text{*\noindent
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   131
The proof is canonical, provided we first show the following simplification
11456
7eb63f63e6c6 revisions and indexing
paulson
parents: 10978
diff changeset
   132
lemma, which also helps to understand what @{term"normif"} does:
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   133
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   134
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   135
lemma [simp]:
10795
9e888d60d3e5 minor edits to Chapters 1-3
paulson
parents: 10171
diff changeset
   136
  "\<forall>t e. valif (normif b t e) env = valif (IF b t e) env";
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   137
(*<*)
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   138
apply(induct_tac b);
9458
c613cd06d5cf apply. -> by
nipkow
parents: 8771
diff changeset
   139
by(auto);
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   140
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   141
theorem "valif (norm b) env = valif b env";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   142
apply(induct_tac b);
9458
c613cd06d5cf apply. -> by
nipkow
parents: 8771
diff changeset
   143
by(auto);
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   144
(*>*)
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   145
text{*\noindent
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   146
Note that the lemma does not have a name, but is implicitly used in the proof
9792
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9541
diff changeset
   147
of the theorem shown above because of the @{text"[simp]"} attribute.
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   148
9792
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9541
diff changeset
   149
But how can we be sure that @{term"norm"} really produces a normal form in
11456
7eb63f63e6c6 revisions and indexing
paulson
parents: 10978
diff changeset
   150
the above sense? We define a function that tests If-expressions for normality:
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   151
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   152
10795
9e888d60d3e5 minor edits to Chapters 1-3
paulson
parents: 10171
diff changeset
   153
consts normal :: "ifex \<Rightarrow> bool";
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   154
primrec
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   155
"normal(CIF b) = True"
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   156
"normal(VIF x) = True"
10795
9e888d60d3e5 minor edits to Chapters 1-3
paulson
parents: 10171
diff changeset
   157
"normal(IF b t e) = (normal t \<and> normal e \<and>
9e888d60d3e5 minor edits to Chapters 1-3
paulson
parents: 10171
diff changeset
   158
     (case b of CIF b \<Rightarrow> True | VIF x \<Rightarrow> True | IF x y z \<Rightarrow> False))";
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   159
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   160
text{*\noindent
11456
7eb63f63e6c6 revisions and indexing
paulson
parents: 10978
diff changeset
   161
Now we prove @{term"normal(norm b)"}. Of course, this requires a lemma about
9792
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9541
diff changeset
   162
normality of @{term"normif"}:
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   163
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   164
12327
5a4d78204492 *** empty log message ***
nipkow
parents: 11456
diff changeset
   165
lemma [simp]: "\<forall>t e. normal(normif b t e) = (normal t \<and> normal e)";
8771
026f37a86ea7 *** empty log message ***
nipkow
parents: 8745
diff changeset
   166
(*<*)
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   167
apply(induct_tac b);
9458
c613cd06d5cf apply. -> by
nipkow
parents: 8771
diff changeset
   168
by(auto);
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   169
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   170
theorem "normal(norm b)";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   171
apply(induct_tac b);
9458
c613cd06d5cf apply. -> by
nipkow
parents: 8771
diff changeset
   172
by(auto);
9933
9feb1e0c4cb3 *** empty log message ***
nipkow
parents: 9792
diff changeset
   173
(*>*)
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   174
9933
9feb1e0c4cb3 *** empty log message ***
nipkow
parents: 9792
diff changeset
   175
text{*\medskip
10795
9e888d60d3e5 minor edits to Chapters 1-3
paulson
parents: 10171
diff changeset
   176
How do we come up with the required lemmas? Try to prove the main theorems
9e888d60d3e5 minor edits to Chapters 1-3
paulson
parents: 10171
diff changeset
   177
without them and study carefully what @{text auto} leaves unproved. This 
9e888d60d3e5 minor edits to Chapters 1-3
paulson
parents: 10171
diff changeset
   178
can provide the clue.  The necessity of universal quantification
9933
9feb1e0c4cb3 *** empty log message ***
nipkow
parents: 9792
diff changeset
   179
(@{text"\<forall>t e"}) in the two lemmas is explained in
9feb1e0c4cb3 *** empty log message ***
nipkow
parents: 9792
diff changeset
   180
\S\ref{sec:InductionHeuristics}
9feb1e0c4cb3 *** empty log message ***
nipkow
parents: 9792
diff changeset
   181
9feb1e0c4cb3 *** empty log message ***
nipkow
parents: 9792
diff changeset
   182
\begin{exercise}
9feb1e0c4cb3 *** empty log message ***
nipkow
parents: 9792
diff changeset
   183
  We strengthen the definition of a @{term normal} If-expression as follows:
9feb1e0c4cb3 *** empty log message ***
nipkow
parents: 9792
diff changeset
   184
  the first argument of all @{term IF}s must be a variable. Adapt the above
9feb1e0c4cb3 *** empty log message ***
nipkow
parents: 9792
diff changeset
   185
  development to this changed requirement. (Hint: you may need to formulate
9feb1e0c4cb3 *** empty log message ***
nipkow
parents: 9792
diff changeset
   186
  some of the goals as implications (@{text"\<longrightarrow>"}) rather than
9feb1e0c4cb3 *** empty log message ***
nipkow
parents: 9792
diff changeset
   187
  equalities (@{text"="}).)
9feb1e0c4cb3 *** empty log message ***
nipkow
parents: 9792
diff changeset
   188
\end{exercise}
11456
7eb63f63e6c6 revisions and indexing
paulson
parents: 10978
diff changeset
   189
\index{boolean expressions example|)}
9933
9feb1e0c4cb3 *** empty log message ***
nipkow
parents: 9792
diff changeset
   190
*}
9feb1e0c4cb3 *** empty log message ***
nipkow
parents: 9792
diff changeset
   191
(*<*)
8771
026f37a86ea7 *** empty log message ***
nipkow
parents: 8745
diff changeset
   192
end
026f37a86ea7 *** empty log message ***
nipkow
parents: 8745
diff changeset
   193
(*>*)