src/HOL/UNITY/Extend.thy
author paulson
Fri, 21 May 1999 10:58:47 +0200
changeset 6677 629b4b3d5bee
parent 6297 5b9fbdfe22b7
child 7342 532841541d73
permissions -rw-r--r--
made definition more readable
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
6297
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
     1
(*  Title:      HOL/UNITY/Extend.thy
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
     2
    ID:         $Id$
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
     4
    Copyright   1998  University of Cambridge
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
     5
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
     6
Extending of state sets
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
     7
  function f (forget)    maps the extended state to the original state
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
     8
  function g (forgotten) maps the extended state to the "extending part"
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
     9
*)
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    10
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    11
Extend = Union + Comp +
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    12
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    13
constdefs
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    14
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    15
  extend_set :: "['a*'b => 'c, 'a set] => 'c set"
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    16
    "extend_set h A == h `` (A Times UNIV)"
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    17
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    18
  extend_act :: "['a*'b => 'c, ('a*'a) set] => ('c * 'c) set"
6677
629b4b3d5bee made definition more readable
paulson
parents: 6297
diff changeset
    19
    "extend_act h act == UN (s,s'): act. UN y. {(h(s,y), h(s',y))}"
6297
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    20
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    21
  extend :: "['a*'b => 'c, 'a program] => 'c program"
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    22
    "extend h F == mk_program (extend_set h (Init F),
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    23
			       extend_act h `` Acts F)"
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    24
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    25
locale Extend =
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    26
  fixes 
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    27
    f       :: 'c => 'a
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    28
    g       :: 'c => 'b
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    29
    h       :: "'a*'b => 'c"    (*isomorphism between 'a * 'b and 'c *)
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    30
    slice   :: ['c set, 'b] => 'a set
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    31
    f_act   :: "('c * 'c) set => ('a*'a) set"
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    32
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    33
  assumes
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    34
    inj_h  "inj h"
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    35
    surj_h "surj h"
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    36
  defines
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    37
    f_def       "f z == fst (inv h z)"
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    38
    g_def       "g z == snd (inv h z)"
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    39
    slice_def   "slice Z y == {x. h(x,y) : Z}"
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    40
    f_act_def   "f_act act == (%(z,z'). (f z, f z')) `` act"
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    41
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    42
end