| author | wenzelm | 
| Sat, 14 Dec 2024 22:04:39 +0100 | |
| changeset 81592 | 775dc5903ed5 | 
| parent 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 | 
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
58889diff
changeset | 11 | | 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: 
58889diff
changeset | 12 | | 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: 
58889diff
changeset | 13 | | 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: 
58889diff
changeset | 14 | | While bexp com (\<open>(WHILE _/ DO _)\<close> [0, 61] 61) | 
| 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
58889diff
changeset | 15 |       | Var    vname com        (\<open>(1{VAR _;/ _})\<close>)
 | 
| 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
58889diff
changeset | 16 |       | Proc   pname com com    (\<open>(1{PROC _ = _;/ _})\<close>)
 | 
| 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: 
51019diff
changeset | 20 | {VAR ''x'';
 | 
| 
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
 nipkow parents: 
51019diff
changeset | 21 |  {PROC ''p'' = ''x'' ::= N 1;
 | 
| 
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
 nipkow parents: 
51019diff
changeset | 22 |   {PROC ''q'' = CALL ''p'';
 | 
| 
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
 nipkow parents: 
51019diff
changeset | 23 |    {VAR ''x'';
 | 
| 
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
 nipkow parents: 
51019diff
changeset | 24 | ''x'' ::= N 2;; | 
| 
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
 nipkow parents: 
51019diff
changeset | 25 |     {PROC ''p'' = ''x'' ::= N 3;
 | 
| 
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
 nipkow parents: 
51019diff
changeset | 26 | CALL ''q'';; ''y'' ::= V ''x''}}}}}" | 
| 43158 | 27 | |
| 28 | end |