doc-src/TutorialI/Datatype/ABexpr.thy
author paulson
Fri, 18 May 2001 17:18:43 +0200
changeset 11309 d666f11ca2d4
parent 10971 6852682eaf16
child 11458 09a6c44a48ea
permissions -rw-r--r--
minor suggestions by Tanja Vos
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 ABexpr = 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
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     5
text{*
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     6
Sometimes it is necessary to define two datatypes that depend on each
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     7
other. This is called \textbf{mutual recursion}. As an example consider a
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     8
language of arithmetic and boolean expressions where
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     9
\begin{itemize}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    10
\item arithmetic expressions contain boolean expressions because there are
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    11
  conditional arithmetic expressions like ``if $m<n$ then $n-m$ else $m-n$'',
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    12
  and
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    13
\item boolean expressions contain arithmetic expressions because of
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    14
  comparisons like ``$m<n$''.
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    15
\end{itemize}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    16
In Isabelle this becomes
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    17
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    18
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    19
datatype 'a aexp = IF   "'a bexp" "'a aexp" "'a aexp"
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    20
                 | Sum  "'a aexp" "'a aexp"
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    21
                 | Diff "'a aexp" "'a aexp"
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    22
                 | Var 'a
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    23
                 | Num nat
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    24
and      'a bexp = Less "'a aexp" "'a aexp"
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    25
                 | And  "'a bexp" "'a bexp"
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    26
                 | Neg  "'a bexp";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    27
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    28
text{*\noindent
9792
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9689
diff changeset
    29
Type @{text"aexp"} is similar to @{text"expr"} in \S\ref{sec:ExprCompiler},
11309
d666f11ca2d4 minor suggestions by Tanja Vos
paulson
parents: 10971
diff changeset
    30
except that we have added an @{text IF} constructor,
d666f11ca2d4 minor suggestions by Tanja Vos
paulson
parents: 10971
diff changeset
    31
fixed the values to be of type @{typ"nat"} and declared the two binary
d666f11ca2d4 minor suggestions by Tanja Vos
paulson
parents: 10971
diff changeset
    32
operations @{text Sum} and @{term"Diff"}.  Boolean
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    33
expressions can be arithmetic comparisons, conjunctions and negations.
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    34
The semantics is fixed via two evaluation functions
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    35
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    36
10971
6852682eaf16 *** empty log message ***
nipkow
parents: 10171
diff changeset
    37
consts  evala :: "'a aexp \<Rightarrow> ('a \<Rightarrow> nat) \<Rightarrow> nat"
6852682eaf16 *** empty log message ***
nipkow
parents: 10171
diff changeset
    38
        evalb :: "'a bexp \<Rightarrow> ('a \<Rightarrow> nat) \<Rightarrow> bool";
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    39
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    40
text{*\noindent
9792
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9689
diff changeset
    41
that take an expression and an environment (a mapping from variables @{typ"'a"} to values
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9689
diff changeset
    42
@{typ"nat"}) and return its arithmetic/boolean
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    43
value. Since the datatypes are mutually recursive, so are functions that
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    44
operate on them. Hence they need to be defined in a single \isacommand{primrec}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    45
section:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    46
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    47
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    48
primrec
8771
026f37a86ea7 *** empty log message ***
nipkow
parents: 8745
diff changeset
    49
  "evala (IF b a1 a2) env =
026f37a86ea7 *** empty log message ***
nipkow
parents: 8745
diff changeset
    50
     (if evalb b env then evala a1 env else evala a2 env)"
026f37a86ea7 *** empty log message ***
nipkow
parents: 8745
diff changeset
    51
  "evala (Sum a1 a2) env = evala a1 env + evala a2 env"
026f37a86ea7 *** empty log message ***
nipkow
parents: 8745
diff changeset
    52
  "evala (Diff a1 a2) env = evala a1 env - evala a2 env"
026f37a86ea7 *** empty log message ***
nipkow
parents: 8745
diff changeset
    53
  "evala (Var v) env = env v"
026f37a86ea7 *** empty log message ***
nipkow
parents: 8745
diff changeset
    54
  "evala (Num n) env = n"
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    55
8771
026f37a86ea7 *** empty log message ***
nipkow
parents: 8745
diff changeset
    56
  "evalb (Less a1 a2) env = (evala a1 env < evala a2 env)"
10971
6852682eaf16 *** empty log message ***
nipkow
parents: 10171
diff changeset
    57
  "evalb (And b1 b2) env = (evalb b1 env \<and> evalb b2 env)"
6852682eaf16 *** empty log message ***
nipkow
parents: 10171
diff changeset
    58
  "evalb (Neg b) env = (\<not> evalb b env)"
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
text{*\noindent
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    61
In the same fashion we also define two functions that perform substitution:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    62
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    63
10971
6852682eaf16 *** empty log message ***
nipkow
parents: 10171
diff changeset
    64
consts substa :: "('a \<Rightarrow> 'b aexp) \<Rightarrow> 'a aexp \<Rightarrow> 'b aexp"
6852682eaf16 *** empty log message ***
nipkow
parents: 10171
diff changeset
    65
       substb :: "('a \<Rightarrow> 'b aexp) \<Rightarrow> 'a bexp \<Rightarrow> 'b bexp";
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    66
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    67
text{*\noindent
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    68
The first argument is a function mapping variables to expressions, the
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    69
substitution. It is applied to all variables in the second argument. As a
9792
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9689
diff changeset
    70
result, the type of variables in the expression may change from @{typ"'a"}
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9689
diff changeset
    71
to @{typ"'b"}. Note that there are only arithmetic and no boolean variables.
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    72
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    73
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    74
primrec
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    75
  "substa s (IF b a1 a2) =
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    76
     IF (substb s b) (substa s a1) (substa s a2)"
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    77
  "substa s (Sum a1 a2) = Sum (substa s a1) (substa s a2)"
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    78
  "substa s (Diff a1 a2) = Diff (substa s a1) (substa s a2)"
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    79
  "substa s (Var v) = s v"
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    80
  "substa s (Num n) = Num n"
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    81
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    82
  "substb s (Less a1 a2) = Less (substa s a1) (substa s a2)"
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    83
  "substb s (And b1 b2) = And (substb s b1) (substb s b2)"
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    84
  "substb s (Neg b) = Neg (substb s b)";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    85
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    86
text{*
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    87
Now we can prove a fundamental theorem about the interaction between
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    88
evaluation and substitution: applying a substitution $s$ to an expression $a$
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    89
and evaluating the result in an environment $env$ yields the same result as
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    90
evaluation $a$ in the environment that maps every variable $x$ to the value
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    91
of $s(x)$ under $env$. If you try to prove this separately for arithmetic or
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    92
boolean expressions (by induction), you find that you always need the other
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    93
theorem in the induction step. Therefore you need to state and prove both
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    94
theorems simultaneously:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    95
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    96
10971
6852682eaf16 *** empty log message ***
nipkow
parents: 10171
diff changeset
    97
lemma "evala (substa s a) env = evala a (\<lambda>x. evala (s x) env) \<and>
6852682eaf16 *** empty log message ***
nipkow
parents: 10171
diff changeset
    98
        evalb (substb s b) env = evalb b (\<lambda>x. evala (s x) env)";
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    99
apply(induct_tac a and b);
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   100
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   101
txt{*\noindent
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   102
The resulting 8 goals (one for each constructor) are proved in one fell swoop:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   103
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   104
10171
59d6633835fa *** empty log message ***
nipkow
parents: 9792
diff changeset
   105
apply simp_all;
59d6633835fa *** empty log message ***
nipkow
parents: 9792
diff changeset
   106
(*<*)done(*>*)
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   107
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   108
text{*
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   109
In general, given $n$ mutually recursive datatypes $\tau@1$, \dots, $\tau@n$,
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   110
an inductive proof expects a goal of the form
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   111
\[ P@1(x@1)\ \land \dots \land P@n(x@n) \]
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   112
where each variable $x@i$ is of type $\tau@i$. Induction is started by
9792
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9689
diff changeset
   113
\begin{isabelle}
10971
6852682eaf16 *** empty log message ***
nipkow
parents: 10171
diff changeset
   114
\isacommand{apply}@{text"(induct_tac"} $x@1$ \isacommand{and} \dots\ \isacommand{and} $x@n$@{text")"}
9792
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9689
diff changeset
   115
\end{isabelle}
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   116
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   117
\begin{exercise}
9792
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9689
diff changeset
   118
  Define a function @{text"norma"} of type @{typ"'a aexp => 'a aexp"} that
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9689
diff changeset
   119
  replaces @{term"IF"}s with complex boolean conditions by nested
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9689
diff changeset
   120
  @{term"IF"}s where each condition is a @{term"Less"} --- @{term"And"} and
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9689
diff changeset
   121
  @{term"Neg"} should be eliminated completely. Prove that @{text"norma"}
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9689
diff changeset
   122
  preserves the value of an expression and that the result of @{text"norma"}
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9689
diff changeset
   123
  is really normal, i.e.\ no more @{term"And"}s and @{term"Neg"}s occur in
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   124
  it.  ({\em Hint:} proceed as in \S\ref{sec:boolex}).
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   125
\end{exercise}
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
end
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   129
(*>*)