IMP/Denotation.thy
changeset 132 47be9d22a0d6
parent 131 41bf53133ba6
child 137 bca8203b0e91
--- a/IMP/Denotation.thy	Wed Aug 31 15:15:54 1994 +0200
+++ b/IMP/Denotation.thy	Wed Aug 31 16:25:19 1994 +0200
@@ -1,4 +1,4 @@
-(*  Title: 	ZF/IMP/Denotation.thy
+(*  Title: 	HOL/IMP/Denotation.thy
     ID:         $Id$
     Author: 	Heiko Loetzbeyer & Robert Sandner, TUM
     Copyright   1994 TUM
@@ -9,26 +9,26 @@
 Denotation = Com + 
 
 consts
-  A     :: "aexp => env => nat"
-  B     :: "bexp => env => bool"
-  C     :: "com => (env*env)set"
-  Gamma :: "[bexp,com,(env*env)set] => (env*env)set"
+  A     :: "aexp => state => nat"
+  B     :: "bexp => state => bool"
+  C     :: "com => (state*state)set"
+  Gamma :: "[bexp,com,(state*state)set] => (state*state)set"
 
 rules
-  A_nat_def	"A(N(n)) == (%sigma. n)"
-  A_loc_def	"A(X(x)) == (%sigma. sigma(x))" 
-  A_op1_def	"A(Op1(f,a)) == (%sigma. f(A(a,sigma)))"
-  A_op2_def	"A(Op2(f,a0,a1)) == (%sigma. f(A(a0,sigma),A(a1,sigma)))"
+  A_nat_def	"A(N(n)) == (%s. n)"
+  A_loc_def	"A(X(x)) == (%s. s(x))" 
+  A_op1_def	"A(Op1(f,a)) == (%s. f(A(a,s)))"
+  A_op2_def	"A(Op2(f,a0,a1)) == (%s. f(A(a0,s),A(a1,s)))"
 
 
-  B_true_def	"B(true) == (%sigma. True)"
-  B_false_def	"B(false) == (%sigma. False)"
-  B_op_def	"B(ROp(f,a0,a1)) == (%sigma. f(A(a0,sigma),A(a1,sigma)))" 
+  B_true_def	"B(true) == (%s. True)"
+  B_false_def	"B(false) == (%s. False)"
+  B_op_def	"B(ROp(f,a0,a1)) == (%s. f(A(a0,s),A(a1,s)))" 
 
 
-  B_not_def	"B(noti(b)) == (%sigma. ~B(b,sigma))"
-  B_and_def	"B(b0 andi b1) == (%sigma. B(b0,sigma) & B(b1,sigma))"
-  B_or_def	"B(b0 ori b1) == (%sigma. B(b0,sigma) | B(b1,sigma))"
+  B_not_def	"B(noti(b)) == (%s. ~B(b,s))"
+  B_and_def	"B(b0 andi b1) == (%s. B(b0,s) & B(b1,s))"
+  B_or_def	"B(b0 ori b1) == (%s. B(b0,s) | B(b1,s))"
 
   C_skip_def	"C(skip) == id"
   C_assign_def	"C(x := a) == {io . snd(io) = fst(io)[A(a,fst(io))/x]}"