src/HOL/IMP/Denotation.thy
author clasohm
Fri, 03 Mar 1995 12:04:16 +0100
changeset 924 806721cfbf46
child 1151 c820b3cc3df0
permissions -rw-r--r--
new version of HOL/IMP with curried function application
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
924
806721cfbf46 new version of HOL/IMP with curried function application
clasohm
parents:
diff changeset
     1
(*  Title: 	HOL/IMP/Denotation.thy
806721cfbf46 new version of HOL/IMP with curried function application
clasohm
parents:
diff changeset
     2
    ID:         $Id$
806721cfbf46 new version of HOL/IMP with curried function application
clasohm
parents:
diff changeset
     3
    Author: 	Heiko Loetzbeyer & Robert Sandner, TUM
806721cfbf46 new version of HOL/IMP with curried function application
clasohm
parents:
diff changeset
     4
    Copyright   1994 TUM
806721cfbf46 new version of HOL/IMP with curried function application
clasohm
parents:
diff changeset
     5
806721cfbf46 new version of HOL/IMP with curried function application
clasohm
parents:
diff changeset
     6
Denotational semantics of expressions & commands
806721cfbf46 new version of HOL/IMP with curried function application
clasohm
parents:
diff changeset
     7
*)
806721cfbf46 new version of HOL/IMP with curried function application
clasohm
parents:
diff changeset
     8
806721cfbf46 new version of HOL/IMP with curried function application
clasohm
parents:
diff changeset
     9
Denotation = Com + 
806721cfbf46 new version of HOL/IMP with curried function application
clasohm
parents:
diff changeset
    10
806721cfbf46 new version of HOL/IMP with curried function application
clasohm
parents:
diff changeset
    11
types com_den = "(state*state)set"
806721cfbf46 new version of HOL/IMP with curried function application
clasohm
parents:
diff changeset
    12
consts
806721cfbf46 new version of HOL/IMP with curried function application
clasohm
parents:
diff changeset
    13
  A     :: "aexp => state => nat"
806721cfbf46 new version of HOL/IMP with curried function application
clasohm
parents:
diff changeset
    14
  B     :: "bexp => state => bool"
806721cfbf46 new version of HOL/IMP with curried function application
clasohm
parents:
diff changeset
    15
  C     :: "com => com_den"
806721cfbf46 new version of HOL/IMP with curried function application
clasohm
parents:
diff changeset
    16
  Gamma :: "[bexp,com_den] => (com_den => com_den)"
806721cfbf46 new version of HOL/IMP with curried function application
clasohm
parents:
diff changeset
    17
806721cfbf46 new version of HOL/IMP with curried function application
clasohm
parents:
diff changeset
    18
primrec A aexp
806721cfbf46 new version of HOL/IMP with curried function application
clasohm
parents:
diff changeset
    19
  A_nat	"A(N(n)) = (%s. n)"
806721cfbf46 new version of HOL/IMP with curried function application
clasohm
parents:
diff changeset
    20
  A_loc	"A(X(x)) = (%s. s(x))" 
806721cfbf46 new version of HOL/IMP with curried function application
clasohm
parents:
diff changeset
    21
  A_op1	"A(Op1 f a) = (%s. f(A a s))"
806721cfbf46 new version of HOL/IMP with curried function application
clasohm
parents:
diff changeset
    22
  A_op2	"A(Op2 f a0 a1) = (%s. f (A a0 s) (A a1 s))"
806721cfbf46 new version of HOL/IMP with curried function application
clasohm
parents:
diff changeset
    23
806721cfbf46 new version of HOL/IMP with curried function application
clasohm
parents:
diff changeset
    24
primrec B bexp
806721cfbf46 new version of HOL/IMP with curried function application
clasohm
parents:
diff changeset
    25
  B_true  "B(true) = (%s. True)"
806721cfbf46 new version of HOL/IMP with curried function application
clasohm
parents:
diff changeset
    26
  B_false "B(false) = (%s. False)"
806721cfbf46 new version of HOL/IMP with curried function application
clasohm
parents:
diff changeset
    27
  B_op	  "B(ROp f a0 a1) = (%s. f (A a0 s) (A a1 s))" 
806721cfbf46 new version of HOL/IMP with curried function application
clasohm
parents:
diff changeset
    28
  B_not	  "B(noti(b)) = (%s. ~(B b s))"
806721cfbf46 new version of HOL/IMP with curried function application
clasohm
parents:
diff changeset
    29
  B_and	  "B(b0 andi b1) = (%s. (B b0 s) & (B b1 s))"
806721cfbf46 new version of HOL/IMP with curried function application
clasohm
parents:
diff changeset
    30
  B_or	  "B(b0 ori b1) = (%s. (B b0 s) | (B b1 s))"
806721cfbf46 new version of HOL/IMP with curried function application
clasohm
parents:
diff changeset
    31
806721cfbf46 new version of HOL/IMP with curried function application
clasohm
parents:
diff changeset
    32
defs
806721cfbf46 new version of HOL/IMP with curried function application
clasohm
parents:
diff changeset
    33
  Gamma_def	"Gamma b cd ==   \
806721cfbf46 new version of HOL/IMP with curried function application
clasohm
parents:
diff changeset
    34
\		   (%phi.{st. st : (phi O cd) & B b (fst st)} Un \
806721cfbf46 new version of HOL/IMP with curried function application
clasohm
parents:
diff changeset
    35
\	                 {st. st : id & ~B b (fst st)})"
806721cfbf46 new version of HOL/IMP with curried function application
clasohm
parents:
diff changeset
    36
806721cfbf46 new version of HOL/IMP with curried function application
clasohm
parents:
diff changeset
    37
primrec C com
806721cfbf46 new version of HOL/IMP with curried function application
clasohm
parents:
diff changeset
    38
  C_skip    "C(skip) = id"
806721cfbf46 new version of HOL/IMP with curried function application
clasohm
parents:
diff changeset
    39
  C_assign  "C(x := a) = {st . snd(st) = fst(st)[A a (fst st)/x]}"
806721cfbf46 new version of HOL/IMP with curried function application
clasohm
parents:
diff changeset
    40
  C_comp    "C(c0 ; c1) = C(c1) O C(c0)"
806721cfbf46 new version of HOL/IMP with curried function application
clasohm
parents:
diff changeset
    41
  C_if	    "C(ifc b then c0 else c1) =   \
806721cfbf46 new version of HOL/IMP with curried function application
clasohm
parents:
diff changeset
    42
\		   {st. st:C(c0) & (B b (fst st))} Un \
806721cfbf46 new version of HOL/IMP with curried function application
clasohm
parents:
diff changeset
    43
\	           {st. st:C(c1) & ~(B b (fst st))}"
806721cfbf46 new version of HOL/IMP with curried function application
clasohm
parents:
diff changeset
    44
  C_while   "C(while b do c) = lfp (Gamma b (C c))"
806721cfbf46 new version of HOL/IMP with curried function application
clasohm
parents:
diff changeset
    45
806721cfbf46 new version of HOL/IMP with curried function application
clasohm
parents:
diff changeset
    46
end