| 43158 |      1 | header "Extensions and Variations of IMP"
 | 
|  |      2 | 
 | 
|  |      3 | theory Procs imports BExp begin
 | 
|  |      4 | 
 | 
|  |      5 | subsection "Procedures and Local Variables"
 | 
|  |      6 | 
 | 
|  |      7 | datatype
 | 
|  |      8 |   com = SKIP 
 | 
|  |      9 |       | Assign name aexp         ("_ ::= _" [1000, 61] 61)
 | 
|  |     10 |       | Semi   com  com          ("_;/ _"  [60, 61] 60)
 | 
|  |     11 |       | If     bexp com com     ("(IF _/ THEN _/ ELSE _)"  [0, 0, 61] 61)
 | 
|  |     12 |       | While  bexp com         ("(WHILE _/ DO _)"  [0, 61] 61)
 | 
|  |     13 |       | Var    name com         ("(1{VAR _;;/ _})")
 | 
|  |     14 |       | Proc   name com com     ("(1{PROC _ = _;;/ _})")
 | 
|  |     15 |       | CALL   name
 | 
|  |     16 | 
 | 
|  |     17 | definition "test_com =
 | 
|  |     18 | {VAR ''x'';;
 | 
|  |     19 |  ''x'' ::= N 0;
 | 
|  |     20 |  {PROC ''p'' = ''x'' ::= Plus (V ''x'') (V ''x'');;
 | 
|  |     21 |   {PROC ''q'' = CALL ''p'';;
 | 
|  |     22 |    {VAR ''x'';;
 | 
|  |     23 |     ''x'' ::= N 5;
 | 
|  |     24 |     {PROC ''p'' = ''x'' ::= Plus (V ''x'') (N 1);;
 | 
|  |     25 |      CALL ''q''; ''y'' ::= V ''x''}}}}}"
 | 
|  |     26 | 
 | 
|  |     27 | end
 |