author | oheimb |
Thu, 12 Mar 1998 13:17:13 +0100 | |
changeset 4743 | b3bfcbd9fb93 |
parent 1900 | c7a869229091 |
child 5102 | 8c782c25a11e |
permissions | -rw-r--r-- |
1697
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
1 |
(* Title: HOL/IMP/Expr.thy |
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
2 |
ID: $Id$ |
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
3 |
Author: Heiko Loetzbeyer & Robert Sandner & Tobias Nipkow, TUM |
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
4 |
Copyright 1994 TUM |
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
5 |
|
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
6 |
Arithmetic expressions and Boolean expressions. |
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
7 |
Not used in the rest of the language, but included for completeness. |
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
8 |
*) |
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
9 |
|
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
10 |
Expr = Arith + |
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
11 |
|
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
12 |
(** Arithmetic expressions **) |
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
13 |
types loc |
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
14 |
state = "loc => nat" |
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
15 |
n2n = "nat => nat" |
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
16 |
n2n2n = "nat => nat => nat" |
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
17 |
|
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
18 |
arities loc :: term |
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
19 |
|
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
20 |
datatype |
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
21 |
aexp = N nat |
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
22 |
| X loc |
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
23 |
| Op1 n2n aexp |
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
24 |
| Op2 n2n2n aexp aexp |
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
25 |
|
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
26 |
(** Evaluation of arithmetic expressions **) |
1729 | 27 |
consts evala :: "((aexp*state) * nat) set" |
28 |
"-a->" :: "[aexp*state,nat] => bool" (infixl 50) |
|
1697
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
29 |
translations |
1729 | 30 |
"aesig -a-> n" == "(aesig,n) : evala" |
1789 | 31 |
inductive evala |
1697
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
32 |
intrs |
1729 | 33 |
N "(N(n),s) -a-> n" |
34 |
X "(X(x),s) -a-> s(x)" |
|
35 |
Op1 "(e,s) -a-> n ==> (Op1 f e,s) -a-> f(n)" |
|
36 |
Op2 "[| (e0,s) -a-> n0; (e1,s) -a-> n1 |] |
|
37 |
==> (Op2 f e0 e1,s) -a-> f n0 n1" |
|
1697
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
38 |
|
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
39 |
types n2n2b = "[nat,nat] => bool" |
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
40 |
|
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
41 |
(** Boolean expressions **) |
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
42 |
|
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
43 |
datatype |
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
44 |
bexp = true |
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
45 |
| false |
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
46 |
| ROp n2n2b aexp aexp |
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
47 |
| noti bexp |
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
48 |
| andi bexp bexp (infixl 60) |
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
49 |
| ori bexp bexp (infixl 60) |
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
50 |
|
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
51 |
(** Evaluation of boolean expressions **) |
1729 | 52 |
consts evalb :: "((bexp*state) * bool)set" |
53 |
"-b->" :: "[bexp*state,bool] => bool" (infixl 50) |
|
1697
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
54 |
|
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
55 |
translations |
1729 | 56 |
"besig -b-> b" == "(besig,b) : evalb" |
1697
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
57 |
|
1789 | 58 |
inductive evalb |
1697
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
59 |
intrs (*avoid clash with ML constructors true, false*) |
1729 | 60 |
tru "(true,s) -b-> True" |
61 |
fls "(false,s) -b-> False" |
|
62 |
ROp "[| (a0,s) -a-> n0; (a1,s) -a-> n1 |] |
|
63 |
==> (ROp f a0 a1,s) -b-> f n0 n1" |
|
64 |
noti "(b,s) -b-> w ==> (noti(b),s) -b-> (~w)" |
|
65 |
andi "[| (b0,s) -b-> w0; (b1,s) -b-> w1 |] |
|
66 |
==> (b0 andi b1,s) -b-> (w0 & w1)" |
|
67 |
ori "[| (b0,s) -b-> w0; (b1,s) -b-> w1 |] |
|
68 |
==> (b0 ori b1,s) -b-> (w0 | w1)" |
|
1697
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
69 |
|
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
70 |
(** Denotational semantics of arithemtic and boolean expressions **) |
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
71 |
consts |
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
72 |
A :: aexp => state => nat |
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
73 |
B :: bexp => state => bool |
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
74 |
|
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
75 |
primrec A aexp |
1900 | 76 |
"A(N(n)) = (%s. n)" |
77 |
"A(X(x)) = (%s. s(x))" |
|
78 |
"A(Op1 f a) = (%s. f(A a s))" |
|
79 |
"A(Op2 f a0 a1) = (%s. f (A a0 s) (A a1 s))" |
|
1697
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
80 |
|
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
81 |
primrec B bexp |
1900 | 82 |
"B(true) = (%s. True)" |
83 |
"B(false) = (%s. False)" |
|
84 |
"B(ROp f a0 a1) = (%s. f (A a0 s) (A a1 s))" |
|
85 |
"B(noti(b)) = (%s. ~(B b s))" |
|
86 |
"B(b0 andi b1) = (%s. (B b0 s) & (B b1 s))" |
|
87 |
"B(b0 ori b1) = (%s. (B b0 s) | (B b1 s))" |
|
1697
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
88 |
|
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
89 |
end |
1900 | 90 |