author | wenzelm |
Wed, 25 Jun 2008 22:01:34 +0200 | |
changeset 27362 | a6dc1769fdda |
parent 23746 | a455e69c31cc |
child 37736 | 2bf3a2cb5e58 |
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 |
|
12431 | 7 |
header "Expressions" |
8 |
||
16417 | 9 |
theory Expr imports Main begin |
1697
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
10 |
|
12431 | 11 |
text {* |
12 |
Arithmetic expressions and Boolean expressions. |
|
13 |
Not used in the rest of the language, but included for completeness. |
|
14 |
*} |
|
1697
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
15 |
|
12431 | 16 |
subsection "Arithmetic expressions" |
12546 | 17 |
typedecl loc |
12431 | 18 |
|
19 |
types |
|
20 |
state = "loc => nat" |
|
1697
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
21 |
|
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
22 |
datatype |
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
23 |
aexp = N nat |
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
24 |
| X loc |
12546 | 25 |
| Op1 "nat => nat" aexp |
26 |
| Op2 "nat => nat => nat" aexp aexp |
|
1697
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
27 |
|
12431 | 28 |
subsection "Evaluation of arithmetic expressions" |
23746 | 29 |
|
30 |
inductive |
|
31 |
evala :: "[aexp*state,nat] => bool" (infixl "-a->" 50) |
|
32 |
where |
|
12431 | 33 |
N: "(N(n),s) -a-> n" |
23746 | 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 |] |
|
12431 | 37 |
==> (Op2 f e0 e1,s) -a-> f n0 n1" |
38 |
||
39 |
lemmas [intro] = N X Op1 Op2 |
|
1697
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 |
|
12431 | 42 |
subsection "Boolean expressions" |
1697
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
43 |
|
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
44 |
datatype |
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
45 |
bexp = true |
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
46 |
| false |
12546 | 47 |
| ROp "nat => nat => bool" aexp aexp |
1697
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
48 |
| noti bexp |
22818 | 49 |
| andi bexp bexp (infixl "andi" 60) |
50 |
| ori bexp bexp (infixl "ori" 60) |
|
1697
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
51 |
|
12431 | 52 |
subsection "Evaluation of boolean expressions" |
1697
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
53 |
|
23746 | 54 |
inductive |
55 |
evalb :: "[bexp*state,bool] => bool" (infixl "-b->" 50) |
|
12431 | 56 |
-- "avoid clash with ML constructors true, false" |
23746 | 57 |
where |
12431 | 58 |
tru: "(true,s) -b-> True" |
23746 | 59 |
| fls: "(false,s) -b-> False" |
60 |
| ROp: "[| (a0,s) -a-> n0; (a1,s) -a-> n1 |] |
|
12431 | 61 |
==> (ROp f a0 a1,s) -b-> f n0 n1" |
23746 | 62 |
| noti: "(b,s) -b-> w ==> (noti(b),s) -b-> (~w)" |
63 |
| andi: "[| (b0,s) -b-> w0; (b1,s) -b-> w1 |] |
|
1729 | 64 |
==> (b0 andi b1,s) -b-> (w0 & w1)" |
23746 | 65 |
| ori: "[| (b0,s) -b-> w0; (b1,s) -b-> w1 |] |
12431 | 66 |
==> (b0 ori b1,s) -b-> (w0 | w1)" |
67 |
||
68 |
lemmas [intro] = tru fls ROp noti andi ori |
|
1697
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
69 |
|
12431 | 70 |
subsection "Denotational semantics of arithmetic and boolean expressions" |
1697
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
71 |
|
27362 | 72 |
primrec A :: "aexp => state => nat" |
73 |
where |
|
1900 | 74 |
"A(N(n)) = (%s. n)" |
27362 | 75 |
| "A(X(x)) = (%s. s(x))" |
76 |
| "A(Op1 f a) = (%s. f(A a s))" |
|
77 |
| "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
|
78 |
|
27362 | 79 |
primrec B :: "bexp => state => bool" |
80 |
where |
|
1900 | 81 |
"B(true) = (%s. True)" |
27362 | 82 |
| "B(false) = (%s. False)" |
83 |
| "B(ROp f a0 a1) = (%s. f (A a0 s) (A a1 s))" |
|
84 |
| "B(noti(b)) = (%s. ~(B b s))" |
|
85 |
| "B(b0 andi b1) = (%s. (B b0 s) & (B b1 s))" |
|
86 |
| "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
|
87 |
|
12431 | 88 |
lemma [simp]: "(N(n),s) -a-> n' = (n = n')" |
89 |
by (rule,cases set: evala) auto |
|
90 |
||
12546 | 91 |
lemma [simp]: "(X(x),sigma) -a-> i = (i = sigma x)" |
12431 | 92 |
by (rule,cases set: evala) auto |
93 |
||
12546 | 94 |
lemma [simp]: |
95 |
"(Op1 f e,sigma) -a-> i = (\<exists>n. i = f n \<and> (e,sigma) -a-> n)" |
|
12431 | 96 |
by (rule,cases set: evala) auto |
97 |
||
98 |
lemma [simp]: |
|
12546 | 99 |
"(Op2 f a1 a2,sigma) -a-> i = |
12431 | 100 |
(\<exists>n0 n1. i = f n0 n1 \<and> (a1, sigma) -a-> n0 \<and> (a2, sigma) -a-> n1)" |
101 |
by (rule,cases set: evala) auto |
|
102 |
||
12546 | 103 |
lemma [simp]: "((true,sigma) -b-> w) = (w=True)" |
12431 | 104 |
by (rule,cases set: evalb) auto |
105 |
||
106 |
lemma [simp]: |
|
12546 | 107 |
"((false,sigma) -b-> w) = (w=False)" |
12431 | 108 |
by (rule,cases set: evalb) auto |
109 |
||
110 |
lemma [simp]: |
|
12546 | 111 |
"((ROp f a0 a1,sigma) -b-> w) = |
112 |
(? m. (a0,sigma) -a-> m & (? n. (a1,sigma) -a-> n & w = f m n))" |
|
23746 | 113 |
by (rule,cases set: evalb) blast+ |
12431 | 114 |
|
12546 | 115 |
lemma [simp]: |
116 |
"((noti(b),sigma) -b-> w) = (? x. (b,sigma) -b-> x & w = (~x))" |
|
23746 | 117 |
by (rule,cases set: evalb) blast+ |
12546 | 118 |
|
119 |
lemma [simp]: |
|
120 |
"((b0 andi b1,sigma) -b-> w) = |
|
121 |
(? x. (b0,sigma) -b-> x & (? y. (b1,sigma) -b-> y & w = (x&y)))" |
|
23746 | 122 |
by (rule,cases set: evalb) blast+ |
12431 | 123 |
|
124 |
lemma [simp]: |
|
12546 | 125 |
"((b0 ori b1,sigma) -b-> w) = |
126 |
(? x. (b0,sigma) -b-> x & (? y. (b1,sigma) -b-> y & w = (x|y)))" |
|
23746 | 127 |
by (rule,cases set: evalb) blast+ |
12431 | 128 |
|
129 |
||
18372 | 130 |
lemma aexp_iff: "((a,s) -a-> n) = (A a s = n)" |
20503 | 131 |
by (induct a arbitrary: n) auto |
12431 | 132 |
|
12546 | 133 |
lemma bexp_iff: |
18372 | 134 |
"((b,s) -b-> w) = (B b s = w)" |
20503 | 135 |
by (induct b arbitrary: w) (auto simp add: aexp_iff) |
12431 | 136 |
|
1697
687f0710c22d
Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff
changeset
|
137 |
end |