| author | wenzelm | 
| Wed, 06 Oct 1999 18:11:37 +0200 | |
| changeset 7754 | 4b1bc1266c8c | 
| parent 7546 | 36b26759147e | 
| child 7826 | c6a8b73b6c2a | 
| permissions | -rw-r--r-- | 
| 6297 | 1  | 
(* Title: HOL/UNITY/Extend.thy  | 
2  | 
ID: $Id$  | 
|
3  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
|
4  | 
Copyright 1998 University of Cambridge  | 
|
5  | 
||
6  | 
Extending of state sets  | 
|
7  | 
function f (forget) maps the extended state to the original state  | 
|
8  | 
function g (forgotten) maps the extended state to the "extending part"  | 
|
9  | 
*)  | 
|
10  | 
||
| 
7399
 
cf780c2bcccf
changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
 
paulson 
parents: 
7378 
diff
changeset
 | 
11  | 
Extend = Guar +  | 
| 6297 | 12  | 
|
13  | 
constdefs  | 
|
14  | 
||
| 7482 | 15  | 
good_map :: "['a*'b => 'c] => bool"  | 
16  | 
"good_map h == surj h & (ALL x y. fst (inv h (h (x,y))) = x)"  | 
|
17  | 
(*Using the locale constant "f", this is f (h (x,y))) = x*)  | 
|
18  | 
||
| 6297 | 19  | 
extend_set :: "['a*'b => 'c, 'a set] => 'c set"  | 
20  | 
"extend_set h A == h `` (A Times UNIV)"  | 
|
21  | 
||
| 7342 | 22  | 
project_set :: "['a*'b => 'c, 'c set] => 'a set"  | 
23  | 
    "project_set h C == {x. EX y. h(x,y) : C}"
 | 
|
24  | 
||
25  | 
  extend_act :: "['a*'b => 'c, ('a*'a) set] => ('c*'c) set"
 | 
|
| 6677 | 26  | 
    "extend_act h act == UN (s,s'): act. UN y. {(h(s,y), h(s',y))}"
 | 
| 6297 | 27  | 
|
| 
7546
 
36b26759147e
project_act no longer has a special case to allow identity actions
 
paulson 
parents: 
7537 
diff
changeset
 | 
28  | 
(*Argument C allows weak safety laws to be projected*)  | 
| 7537 | 29  | 
  project_act :: "['c set, 'a*'b => 'c, ('c*'c) set] => ('a*'a) set"
 | 
30  | 
"project_act C h act ==  | 
|
| 
7546
 
36b26759147e
project_act no longer has a special case to allow identity actions
 
paulson 
parents: 
7537 
diff
changeset
 | 
31  | 
         {(x,x'). EX y y'. (h(x,y), h(x',y')) : act & (h(x,y) : C)}"
 | 
| 7342 | 32  | 
|
| 6297 | 33  | 
extend :: "['a*'b => 'c, 'a program] => 'c program"  | 
34  | 
"extend h F == mk_program (extend_set h (Init F),  | 
|
35  | 
extend_act h `` Acts F)"  | 
|
36  | 
||
| 7537 | 37  | 
project :: "['c set, 'a*'b => 'c, 'c program] => 'a program"  | 
38  | 
"project C h F == mk_program (project_set h (Init F),  | 
|
39  | 
project_act C h `` Acts F)"  | 
|
| 7342 | 40  | 
|
| 6297 | 41  | 
locale Extend =  | 
42  | 
fixes  | 
|
43  | 
f :: 'c => 'a  | 
|
44  | 
g :: 'c => 'b  | 
|
45  | 
h :: "'a*'b => 'c" (*isomorphism between 'a * 'b and 'c *)  | 
|
46  | 
slice :: ['c set, 'b] => 'a set  | 
|
47  | 
||
48  | 
assumes  | 
|
| 7482 | 49  | 
good_h "good_map h"  | 
| 6297 | 50  | 
defines  | 
51  | 
f_def "f z == fst (inv h z)"  | 
|
52  | 
g_def "g z == snd (inv h z)"  | 
|
53  | 
    slice_def   "slice Z y == {x. h(x,y) : Z}"
 | 
|
54  | 
||
55  | 
end  |