| 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 | 
 | 
|  |     11 | Extend = Union + Comp +
 | 
|  |     12 | 
 | 
|  |     13 | constdefs
 | 
|  |     14 | 
 | 
|  |     15 |   extend_set :: "['a*'b => 'c, 'a set] => 'c set"
 | 
|  |     16 |     "extend_set h A == h `` (A Times UNIV)"
 | 
|  |     17 | 
 | 
|  |     18 |   extend_act :: "['a*'b => 'c, ('a*'a) set] => ('c * 'c) set"
 | 
|  |     19 |     "extend_act h == (%act. UN (s,s'): act. UN y. {(h(s,y), h(s',y))})"
 | 
|  |     20 | 
 | 
|  |     21 |   extend :: "['a*'b => 'c, 'a program] => 'c program"
 | 
|  |     22 |     "extend h F == mk_program (extend_set h (Init F),
 | 
|  |     23 | 			       extend_act h `` Acts F)"
 | 
|  |     24 | 
 | 
|  |     25 | locale Extend =
 | 
|  |     26 |   fixes 
 | 
|  |     27 |     f       :: 'c => 'a
 | 
|  |     28 |     g       :: 'c => 'b
 | 
|  |     29 |     h       :: "'a*'b => 'c"    (*isomorphism between 'a * 'b and 'c *)
 | 
|  |     30 |     slice   :: ['c set, 'b] => 'a set
 | 
|  |     31 |     f_act   :: "('c * 'c) set => ('a*'a) set"
 | 
|  |     32 | 
 | 
|  |     33 |   assumes
 | 
|  |     34 |     inj_h  "inj h"
 | 
|  |     35 |     surj_h "surj h"
 | 
|  |     36 |   defines
 | 
|  |     37 |     f_def       "f z == fst (inv h z)"
 | 
|  |     38 |     g_def       "g z == snd (inv h z)"
 | 
|  |     39 |     slice_def   "slice Z y == {x. h(x,y) : Z}"
 | 
|  |     40 |     f_act_def   "f_act act == (%(z,z'). (f z, f z')) `` act"
 | 
|  |     41 | 
 | 
|  |     42 | end
 |