src/HOL/IMP/Com.thy
author wenzelm
Wed, 12 Mar 2025 11:39:00 +0100
changeset 82265 4b875a4c83b0
parent 80914 d97fdabd9e2b
permissions -rw-r--r--
update for release;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
58889
5b7a9633cfa8 modernized header uniformly as section;
wenzelm
parents: 58310
diff changeset
     1
section "IMP --- A Simple Imperative Language"
924
806721cfbf46 new version of HOL/IMP with curried function application
clasohm
parents:
diff changeset
     2
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 42174
diff changeset
     3
theory Com imports BExp begin
924
806721cfbf46 new version of HOL/IMP with curried function application
clasohm
parents:
diff changeset
     4
58310
91ea607a34d8 updated news
blanchet
parents: 58305
diff changeset
     5
datatype
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 42174
diff changeset
     6
  com = SKIP 
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 58889
diff changeset
     7
      | 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
     8
      | 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
     9
      | 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
    10
      | While  bexp com         (\<open>(WHILE _/ DO _)\<close>  [0, 61] 61)
924
806721cfbf46 new version of HOL/IMP with curried function application
clasohm
parents:
diff changeset
    11
806721cfbf46 new version of HOL/IMP with curried function application
clasohm
parents:
diff changeset
    12
end