author | wenzelm |
Sun, 16 Sep 2018 20:33:37 +0200 | |
changeset 69002 | f0d4b1db0ea0 |
parent 58889 | 5b7a9633cfa8 |
child 80914 | d97fdabd9e2b |
permissions | -rw-r--r-- |
58889 | 1 |
section "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 |
|
58310 | 5 |
datatype |
43141 | 6 |
com = SKIP |
45321 | 7 |
| Assign vname aexp ("_ ::= _" [1000, 61] 61) |
52046
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents:
47818
diff
changeset
|
8 |
| Seq com com ("_;;/ _" [60, 61] 60) |
43141 | 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 |