| author | blanchet | 
| Thu, 28 Jul 2011 11:43:45 +0200 | |
| changeset 43996 | 4d1270ddf042 | 
| parent 43141 | 11fce8564415 | 
| child 45212 | e87feee00a4c | 
| permissions | -rw-r--r-- | 
| 43141 | 1 | header "IMP --- A Simple Imperative Language" | 
| 924 
806721cfbf46
new version of HOL/IMP with curried function application
 clasohm parents: diff
changeset | 2 | |
| 43141 | 3 | theory Com imports BExp begin | 
| 924 
806721cfbf46
new version of HOL/IMP with curried function application
 clasohm parents: diff
changeset | 4 | |
| 
806721cfbf46
new version of HOL/IMP with curried function application
 clasohm parents: diff
changeset | 5 | datatype | 
| 43141 | 6 | com = SKIP | 
| 7 |       | Assign name aexp         ("_ ::= _" [1000, 61] 61)
 | |
| 8 |       | Semi   com  com          ("_;/ _"  [60, 61] 60)
 | |
| 9 |       | If     bexp com com     ("(IF _/ THEN _/ ELSE _)"  [0, 0, 61] 61)
 | |
| 10 |       | While  bexp com         ("(WHILE _/ DO _)"  [0, 61] 61)
 | |
| 924 
806721cfbf46
new version of HOL/IMP with curried function application
 clasohm parents: diff
changeset | 11 | |
| 
806721cfbf46
new version of HOL/IMP with curried function application
 clasohm parents: diff
changeset | 12 | end |