author | nipkow |
Mon, 13 May 2002 15:27:28 +0200 | |
changeset 13145 | 59bc43b51aa2 |
parent 10064 | 1a77667b21ef |
child 13790 | 8d7e9fce8c50 |
permissions | -rw-r--r-- |
7630 | 1 |
(* Title: HOL/UNITY/Project.ML |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1999 University of Cambridge |
|
5 |
||
6 |
Projections of state sets (also of actions and programs) |
|
7 |
||
8 |
Inheritance of GUARANTEES properties under extension |
|
9 |
*) |
|
10 |
||
7826
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
paulson
parents:
7689
diff
changeset
|
11 |
Project = Extend + |
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
paulson
parents:
7689
diff
changeset
|
12 |
|
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
paulson
parents:
7689
diff
changeset
|
13 |
constdefs |
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
paulson
parents:
7689
diff
changeset
|
14 |
projecting :: "['c program => 'c set, 'a*'b => 'c, |
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
paulson
parents:
7689
diff
changeset
|
15 |
'a program, 'c program set, 'a program set] => bool" |
10064
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents:
8055
diff
changeset
|
16 |
"projecting C h F X' X == |
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents:
8055
diff
changeset
|
17 |
ALL G. extend h F Join G : X' --> F Join project h (C G) G : X" |
7826
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
paulson
parents:
7689
diff
changeset
|
18 |
|
10064
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents:
8055
diff
changeset
|
19 |
extending :: "['c program => 'c set, 'a*'b => 'c, 'a program, |
8055
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
7947
diff
changeset
|
20 |
'c program set, 'a program set] => bool" |
10064
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents:
8055
diff
changeset
|
21 |
"extending C h F Y' Y == |
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents:
8055
diff
changeset
|
22 |
ALL G. extend h F ok G --> F Join project h (C G) G : Y |
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents:
8055
diff
changeset
|
23 |
--> extend h F Join G : Y'" |
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents:
8055
diff
changeset
|
24 |
|
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents:
8055
diff
changeset
|
25 |
subset_closed :: "'a set set => bool" |
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents:
8055
diff
changeset
|
26 |
"subset_closed U == ALL A: U. Pow A <= U" |
7826
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
paulson
parents:
7689
diff
changeset
|
27 |
|
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
paulson
parents:
7689
diff
changeset
|
28 |
end |