src/HOL/UNITY/Project.thy
changeset 7630 d0e4a6f1f05c
child 7689 affe0c2fdfbf
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/UNITY/Project.thy	Wed Sep 29 13:13:06 1999 +0200
@@ -0,0 +1,22 @@
+(*  Title:      HOL/UNITY/Project.ML
+    ID:         $Id$
+    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
+    Copyright   1999  University of Cambridge
+
+Projections of state sets (also of actions and programs)
+
+Inheritance of GUARANTEES properties under extension
+*)
+
+Project = Extend +
+
+
+constdefs
+
+  restr_act :: "['c set, 'a*'b => 'c, ('a*'a) set] => ('a*'a) set"
+    "restr_act C h act == project_act C h (extend_act h act)"
+
+  restr :: "['c set, 'a*'b => 'c, 'a program] => 'a program"
+    "restr C h F == project C h (extend h F)"
+
+end