| author | paulson | 
| Fri, 20 Aug 2004 12:20:09 +0200 | |
| changeset 15150 | c7af682b9ee5 | 
| parent 12338 | de0f4a63baa5 | 
| child 17477 | ceb42ea2f223 | 
| permissions | -rw-r--r-- | 
| 8177 | 1 | (* Title: HOL/IMPP/Com.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: David von Oheimb (based on a theory by Tobias Nipkow et al), TUM | |
| 4 | Copyright 1999 TUM | |
| 5 | ||
| 6 | Semantics of arithmetic and boolean expressions, Syntax of commands | |
| 7 | *) | |
| 8 | ||
| 9 | Com = Main + | |
| 10 | ||
| 11 | types val = nat (* for the meta theory, this may be anything, but with | |
| 12 | current Isabelle, types cannot be refined later *) | |
| 13 | types glb | |
| 14 | loc | |
| 12338 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
8177diff
changeset | 15 | arities (*val,*)glb,loc :: type | 
| 8177 | 16 | consts Arg, Res :: loc | 
| 17 | ||
| 18 | datatype vname = Glb glb | Loc loc | |
| 19 | types globs = glb => val | |
| 20 | locals = loc => val | |
| 21 | datatype state = st globs locals | |
| 22 | (* for the meta theory, the following would be sufficient: | |
| 23 | types state | |
| 12338 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
8177diff
changeset | 24 | arities state::type | 
| 8177 | 25 | consts st:: [globs , locals] => state | 
| 26 | *) | |
| 27 | types aexp = state => val | |
| 28 | bexp = state => bool | |
| 29 | ||
| 30 | types pname | |
| 12338 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
8177diff
changeset | 31 | arities pname :: type | 
| 8177 | 32 | |
| 33 | datatype com | |
| 34 | = SKIP | |
| 35 |       | Ass   vname aexp	("_:==_"	        [65, 65    ] 60)
 | |
| 36 |       | Local loc aexp com	("LOCAL _:=_ IN _"	[65,  0, 61] 60)
 | |
| 37 |       | Semi  com  com		("_;; _"	        [59, 60    ] 59)
 | |
| 38 |       | Cond  bexp com com	("IF _ THEN _ ELSE _"	[65, 60, 61] 60)
 | |
| 39 |       | While bexp com		("WHILE _ DO _"		[65,     61] 60)
 | |
| 40 | | BODY pname | |
| 41 |       | Call  vname pname aexp	("_:=CALL _'(_')"	[65, 65,  0] 60)
 | |
| 42 | ||
| 43 | consts bodies :: "(pname * com) list"(* finitely many procedure definitions *) | |
| 44 | body :: " pname ~=> com" | |
| 45 | defs body_def "body == map_of bodies" | |
| 46 | ||
| 47 | ||
| 48 | (* Well-typedness: all procedures called must exist *) | |
| 49 | consts WTs :: com set | |
| 50 | syntax WT :: com => bool | |
| 51 | translations "WT c" == "c : WTs" | |
| 52 | ||
| 53 | inductive WTs intrs | |
| 54 | ||
| 55 | Skip "WT SKIP" | |
| 56 | ||
| 57 | Assign "WT (X :== a)" | |
| 58 | ||
| 59 | Local "WT c ==> | |
| 60 | WT (LOCAL Y := a IN c)" | |
| 61 | ||
| 62 | Semi "[| WT c0; WT c1 |] ==> | |
| 63 | WT (c0;; c1)" | |
| 64 | ||
| 65 | If "[| WT c0; WT c1 |] ==> | |
| 66 | WT (IF b THEN c0 ELSE c1)" | |
| 67 | ||
| 68 | While "WT c ==> | |
| 69 | WT (WHILE b DO c)" | |
| 70 | ||
| 71 | Body "body pn ~= None ==> | |
| 72 | WT (BODY pn)" | |
| 73 | ||
| 74 | Call "WT (BODY pn) ==> | |
| 75 | WT (X:=CALL pn(a))" | |
| 76 | ||
| 77 | constdefs | |
| 78 | WT_bodies :: bool | |
| 79 | "WT_bodies == !(pn,b):set bodies. WT b" | |
| 80 | ||
| 81 | end |