| author | paulson <lp15@cam.ac.uk> |
| Thu, 10 Apr 2025 17:07:18 +0100 | |
| changeset 82470 | 785615e37846 |
| parent 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 |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
58889
diff
changeset
|
7 |
| Assign vname aexp (\<open>_ ::= _\<close> [1000, 61] 61) |
|
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
58889
diff
changeset
|
8 |
| Seq com com (\<open>_;;/ _\<close> [60, 61] 60) |
|
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
58889
diff
changeset
|
9 |
| If bexp com com (\<open>(IF _/ THEN _/ ELSE _)\<close> [0, 0, 61] 61) |
|
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
58889
diff
changeset
|
10 |
| While bexp com (\<open>(WHILE _/ DO _)\<close> [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 |