| author | blanchet | 
| Thu, 13 Mar 2014 13:18:14 +0100 | |
| changeset 56092 | 1ba075db8fe4 | 
| parent 55416 | dd7992d4a61a | 
| child 58249 | 180f1b3508ed | 
| 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 | ||
| 55416 | 11 | datatype action = New | Loc nat | Free nat | 
| 17244 | 12 | |
| 55416 | 13 | lemma [cong]: "!!x. x = y ==> case_action a b c x = case_action a b c y" | 
| 17244 | 14 | by simp | 
| 15 | ||
| 6008 | 16 | end |