author | wenzelm |
Thu, 01 Sep 2016 21:28:46 +0200 | |
changeset 63763 | 0f61ea70d384 |
parent 58889 | 5b7a9633cfa8 |
child 80914 | d97fdabd9e2b |
permissions | -rw-r--r-- |
58889 | 1 |
section "Extensions and Variations of IMP" |
43158 | 2 |
|
3 |
theory Procs imports BExp begin |
|
4 |
||
5 |
subsection "Procedures and Local Variables" |
|
6 |
||
45212 | 7 |
type_synonym pname = string |
8 |
||
58310 | 9 |
datatype |
43158 | 10 |
com = SKIP |
45212 | 11 |
| Assign vname aexp ("_ ::= _" [1000, 61] 61) |
52046
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents:
51019
diff
changeset
|
12 |
| Seq com com ("_;;/ _" [60, 61] 60) |
43158 | 13 |
| If bexp com com ("(IF _/ THEN _/ ELSE _)" [0, 0, 61] 61) |
14 |
| While bexp com ("(WHILE _/ DO _)" [0, 61] 61) |
|
52046
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents:
51019
diff
changeset
|
15 |
| Var vname com ("(1{VAR _;/ _})") |
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents:
51019
diff
changeset
|
16 |
| Proc pname com com ("(1{PROC _ = _;/ _})") |
45212 | 17 |
| CALL pname |
43158 | 18 |
|
19 |
definition "test_com = |
|
52046
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents:
51019
diff
changeset
|
20 |
{VAR ''x''; |
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents:
51019
diff
changeset
|
21 |
{PROC ''p'' = ''x'' ::= N 1; |
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents:
51019
diff
changeset
|
22 |
{PROC ''q'' = CALL ''p''; |
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents:
51019
diff
changeset
|
23 |
{VAR ''x''; |
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents:
51019
diff
changeset
|
24 |
''x'' ::= N 2;; |
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents:
51019
diff
changeset
|
25 |
{PROC ''p'' = ''x'' ::= N 3; |
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents:
51019
diff
changeset
|
26 |
CALL ''q'';; ''y'' ::= V ''x''}}}}}" |
43158 | 27 |
|
28 |
end |