author | aspinall |
Fri, 30 Sep 2005 18:18:34 +0200 | |
changeset 17740 | fc385ce6187d |
parent 17244 | 0b2ff9541727 |
child 19740 | 6b38551d0798 |
permissions | -rw-r--r-- |
(* Title: HOLCF/IOA/ABP/Action.thy ID: $Id$ Author: Olaf Müller *) header {* The set of all actions of the system *} theory Action imports Main begin datatype action = New | Loc nat | Free nat lemma [cong]: "!!x. x = y ==> action_case a b c x = action_case a b c y" by simp ML {* use_legacy_bindings (the_context ()) *} end