| author | wenzelm | 
| Wed, 25 Jun 2008 22:01:34 +0200 | |
| changeset 27362 | a6dc1769fdda | 
| parent 16417 | 9bc16273c2d4 | 
| child 41589 | bbd861837ebc | 
| permissions | -rw-r--r-- | 
| 12431 | 1  | 
(* Title: HOL/IMP/Com.thy  | 
2  | 
ID: $Id$  | 
|
3  | 
Author: Heiko Loetzbeyer & Robert Sandner & Tobias Nipkow, TUM  | 
|
4  | 
Isar Version: Gerwin Klein, 2001  | 
|
5  | 
Copyright 1994 TUM  | 
|
| 
924
 
806721cfbf46
new version of HOL/IMP with curried function application
 
clasohm 
parents:  
diff
changeset
 | 
6  | 
*)  | 
| 
 
806721cfbf46
new version of HOL/IMP with curried function application
 
clasohm 
parents:  
diff
changeset
 | 
7  | 
|
| 12431 | 8  | 
header "Syntax of Commands"  | 
9  | 
||
| 16417 | 10  | 
theory Com imports Main begin  | 
| 
924
 
806721cfbf46
new version of HOL/IMP with curried function application
 
clasohm 
parents:  
diff
changeset
 | 
11  | 
|
| 12431 | 12  | 
typedecl loc  | 
13  | 
-- "an unspecified (arbitrary) type of locations  | 
|
14  | 
(adresses/names) for variables"  | 
|
15  | 
||
16  | 
types  | 
|
17  | 
  val   = nat -- "or anything else, @{text nat} used in examples"
 | 
|
18  | 
state = "loc \<Rightarrow> val"  | 
|
19  | 
aexp = "state \<Rightarrow> val"  | 
|
20  | 
bexp = "state \<Rightarrow> bool"  | 
|
21  | 
-- "arithmetic and boolean expressions are not modelled explicitly here,"  | 
|
22  | 
-- "they are just functions on states"  | 
|
| 
924
 
806721cfbf46
new version of HOL/IMP with curried function application
 
clasohm 
parents:  
diff
changeset
 | 
23  | 
|
| 
 
806721cfbf46
new version of HOL/IMP with curried function application
 
clasohm 
parents:  
diff
changeset
 | 
24  | 
datatype  | 
| 12431 | 25  | 
com = SKIP  | 
26  | 
      | Assign loc aexp         ("_ :== _ " 60)
 | 
|
27  | 
      | Semi   com com          ("_; _"  [60, 60] 10)
 | 
|
28  | 
      | Cond   bexp com com     ("IF _ THEN _ ELSE _"  60)
 | 
|
29  | 
      | While  bexp com         ("WHILE _ DO _"  60)
 | 
|
30  | 
||
| 27362 | 31  | 
notation (latex)  | 
32  | 
  SKIP  ("\<SKIP>") and
 | 
|
33  | 
  Cond  ("\<IF> _ \<THEN> _ \<ELSE> _"  60) and
 | 
|
34  | 
  While  ("\<WHILE> _ \<DO> _"  60)
 | 
|
| 
924
 
806721cfbf46
new version of HOL/IMP with curried function application
 
clasohm 
parents:  
diff
changeset
 | 
35  | 
|
| 
 
806721cfbf46
new version of HOL/IMP with curried function application
 
clasohm 
parents:  
diff
changeset
 | 
36  | 
end  |