| author | paulson | 
| Wed, 28 Jun 2000 10:41:16 +0200 | |
| changeset 9161 | cee6d5aee7c8 | 
| parent 5608 | a82a038a3e7a | 
| child 9241 | f961c1fdff50 | 
| permissions | -rw-r--r-- | 
| 1476 | 1 | (* Title: HOL/IMP/Denotation.thy | 
| 924 
806721cfbf46
new version of HOL/IMP with curried function application
 clasohm parents: diff
changeset | 2 | ID: $Id$ | 
| 1476 | 3 | Author: Heiko Loetzbeyer & Robert Sandner, TUM | 
| 924 
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 | |
| 1696 | 6 | Denotational semantics of commands | 
| 924 
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 | |
| 1696 | 9 | Denotation = Natural + | 
| 924 
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" | 
| 1696 | 12 | |
| 13 | constdefs | |
| 1374 | 14 | Gamma :: [bexp,com_den] => (com_den => com_den) | 
| 1696 | 15 |            "Gamma b cd == (%phi.{(s,t). (s,t) : (phi O cd) & b(s)} Un 
 | 
| 2847 | 16 |                                  {(s,t). s=t & ~b(s)})"
 | 
| 1696 | 17 | |
| 18 | consts | |
| 19 | C :: com => com_den | |
| 924 
806721cfbf46
new version of HOL/IMP with curried function application
 clasohm parents: diff
changeset | 20 | |
| 5183 | 21 | primrec | 
| 5608 | 22 | C_skip "C(SKIP) = Id" | 
| 4897 | 23 |   C_assign  "C(x := a) = {(s,t). t = s[x:=a(s)]}"
 | 
| 924 
806721cfbf46
new version of HOL/IMP with curried function application
 clasohm parents: diff
changeset | 24 | C_comp "C(c0 ; c1) = C(c1) O C(c0)" | 
| 1696 | 25 |   C_if      "C(IF b THEN c1 ELSE c2) = {(s,t). (s,t) : C(c1) & b(s)} Un
 | 
| 26 |                                        {(s,t). (s,t) : C(c2) & ~ b(s)}"
 | |
| 1481 | 27 | C_while "C(WHILE b DO c) = lfp (Gamma b (C c))" | 
| 924 
806721cfbf46
new version of HOL/IMP with curried function application
 clasohm parents: diff
changeset | 28 | |
| 
806721cfbf46
new version of HOL/IMP with curried function application
 clasohm parents: diff
changeset | 29 | end |