src/HOL/IMP/Procs.thy
author wenzelm
Wed, 13 Nov 2024 20:14:24 +0100
changeset 81440 b3c0d032ed6e
parent 80914 d97fdabd9e2b
permissions -rw-r--r--
merged
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
58889
5b7a9633cfa8 modernized header uniformly as section;
wenzelm
parents: 58310
diff changeset
     1
section "Extensions and Variations of IMP"
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
     2
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
     3
theory Procs imports BExp begin
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
     4
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
     5
subsection "Procedures and Local Variables"
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
     6
45212
e87feee00a4c renamed name -> vname
nipkow
parents: 43158
diff changeset
     7
type_synonym pname = string
e87feee00a4c renamed name -> vname
nipkow
parents: 43158
diff changeset
     8
58310
91ea607a34d8 updated news
blanchet
parents: 58305
diff changeset
     9
datatype
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    10
  com = SKIP 
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 58889
diff 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: 58889
diff 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: 58889
diff 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: 58889
diff 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: 58889
diff 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: 58889
diff changeset
    16
      | Proc   pname com com    (\<open>(1{PROC _ = _;/ _})\<close>)
45212
e87feee00a4c renamed name -> vname
nipkow
parents: 43158
diff changeset
    17
      | CALL   pname
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    18
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    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
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    27
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    28
end