src/HOL/IMP/Procs.thy
changeset 47818 151d137f1095
parent 45212 e87feee00a4c
child 51019 146f63c3f024
equal deleted inserted replaced
47817:5d2d63f4363e 47818:151d137f1095
     7 type_synonym pname = string
     7 type_synonym pname = string
     8 
     8 
     9 datatype
     9 datatype
    10   com = SKIP 
    10   com = SKIP 
    11       | Assign vname aexp        ("_ ::= _" [1000, 61] 61)
    11       | Assign vname aexp        ("_ ::= _" [1000, 61] 61)
    12       | Semi   com  com          ("_;/ _"  [60, 61] 60)
    12       | Seq    com  com          ("_;/ _"  [60, 61] 60)
    13       | If     bexp com com     ("(IF _/ THEN _/ ELSE _)"  [0, 0, 61] 61)
    13       | If     bexp com com     ("(IF _/ THEN _/ ELSE _)"  [0, 0, 61] 61)
    14       | While  bexp com         ("(WHILE _/ DO _)"  [0, 61] 61)
    14       | While  bexp com         ("(WHILE _/ DO _)"  [0, 61] 61)
    15       | Var    vname com        ("(1{VAR _;;/ _})")
    15       | Var    vname com        ("(1{VAR _;;/ _})")
    16       | Proc   pname com com    ("(1{PROC _ = _;;/ _})")
    16       | Proc   pname com com    ("(1{PROC _ = _;;/ _})")
    17       | CALL   pname
    17       | CALL   pname