| author | haftmann | 
| Fri, 02 Sep 2011 19:29:36 +0200 | |
| changeset 44663 | 3bc39cfe27fe | 
| parent 42151 | 4da4fc77664b | 
| child 55416 | dd7992d4a61a | 
| permissions | -rw-r--r-- | 
| 42151 | 1 | (* Title: HOL/HOLCF/IOA/Storage/Action.thy | 
| 40945 | 2 | Author: Olaf Müller | 
| 6008 | 3 | *) | 
| 4 | ||
| 17244 | 5 | header {* The set of all actions of the system *}
 | 
| 6 | ||
| 7 | theory Action | |
| 8 | imports Main | |
| 9 | begin | |
| 10 | ||
| 11 | datatype action = New | Loc nat | Free nat | |
| 12 | ||
| 13 | lemma [cong]: "!!x. x = y ==> action_case a b c x = action_case a b c y" | |
| 14 | by simp | |
| 15 | ||
| 6008 | 16 | end |