IMP/Denotation.thy
author wenzelm
Wed, 14 Sep 1994 16:05:28 +0200
changeset 142 760641387b20
parent 138 bf044f0db994
child 144 6254f50e5ec9
permissions -rw-r--r--
replaced lookup_const by Sign.const_type;

(*  Title: 	HOL/IMP/Denotation.thy
    ID:         $Id$
    Author: 	Heiko Loetzbeyer & Robert Sandner, TUM
    Copyright   1994 TUM

Denotational semantics of expressions & commands
*)

Denotation = Com + 

consts
  A     :: "aexp => state => nat"
  B     :: "bexp => state => bool"
  C     :: "com => (state*state)set"
  Gamma :: "[bexp,com,(state*state)set] => (state*state)set"

primrec A aexp
  A_nat	"A(N(n)) = (%s. n)"
  A_loc	"A(X(x)) = (%s. s(x))" 
  A_op1	"A(Op1(f,a)) = (%s. f(A(a,s)))"
  A_op2	"A(Op2(f,a0,a1)) = (%s. f(A(a0,s),A(a1,s)))"

primrec B bexp
  B_true  "B(true) = (%s. True)"
  B_false "B(false) = (%s. False)"
  B_op	  "B(ROp(f,a0,a1)) = (%s. f(A(a0,s),A(a1,s)))" 
  B_not	  "B(noti(b)) = (%s. ~B(b,s))"
  B_and	  "B(b0 andi b1) = (%s. B(b0,s) & B(b1,s))"
  B_or	  "B(b0 ori b1) = (%s. B(b0,s) | B(b1,s))"

defs
  Gamma_def	"Gamma(b,c) ==   \
\		   (%phi.{io. io : (phi O C(c)) & B(b,fst(io))} Un \
\	                 {io. io : id & ~B(b,fst(io))})"

primrec C com
  C_skip    "C(skip) = id"
  C_assign  "C(x := a) = {io . snd(io) = fst(io)[A(a,fst(io))/x]}"
  C_comp    "C(c0 ; c1) = C(c1) O C(c0)"
  C_if	    "C(ifc b then c0 else c1) =   \
\		   {io. io:C(c0) & B(b,fst(io))} Un \
\	           {io. io:C(c1) & ~B(b,fst(io))}"
  C_while   "C(while b do c) = lfp(Gamma(b,c))"

end