| author | wenzelm |
| Sat, 25 Sep 1999 13:08:08 +0200 | |
| changeset 7600 | 73f91da46230 |
| parent 5183 | 89f162de39cf |
| child 8029 | 05446a898852 |
| permissions | -rw-r--r-- |
(* Title: HOL/IMP/Com.thy ID: $Id$ Author: Heiko Loetzbeyer & Robert Sandner & Tobias Nipkow, TUM Copyright 1994 TUM Semantics of arithmetic and boolean expressions Syntax of commands *) Com = Datatype + types val loc state = "loc => val" aexp = "state => val" bexp = state => bool arities loc,val :: term datatype com = SKIP | ":=" loc aexp (infixl 60) | Semi com com ("_; _" [60, 60] 10) | Cond bexp com com ("IF _ THEN _ ELSE _" 60) | While bexp com ("WHILE _ DO _" 60) end