src/HOL/UNITY/Extend.thy
author wenzelm
Fri, 09 Nov 2001 00:09:47 +0100
changeset 12114 a8e860c86252
parent 10834 a7897aebbffc
child 13790 8d7e9fce8c50
permissions -rw-r--r--
eliminated old "symbols" syntax, use "xsymbols" instead;
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
7399
cf780c2bcccf changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
paulson
parents: 7378
diff changeset
    11
Extend = Guar +
6297
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
8948
b797cfa3548d restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
paulson
parents: 8703
diff changeset
    15
  (*MOVE to Relation.thy?*)
b797cfa3548d restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
paulson
parents: 8703
diff changeset
    16
  Restrict :: "[ 'a set, ('a*'b) set] => ('a*'b) set"
b797cfa3548d restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
paulson
parents: 8703
diff changeset
    17
    "Restrict A r == r Int (A <*> UNIV)"
b797cfa3548d restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
paulson
parents: 8703
diff changeset
    18
7482
7badd511844d working snapshot
paulson
parents: 7399
diff changeset
    19
  good_map :: "['a*'b => 'c] => bool"
7badd511844d working snapshot
paulson
parents: 7399
diff changeset
    20
    "good_map h == surj h & (ALL x y. fst (inv h (h (x,y))) = x)"
7badd511844d working snapshot
paulson
parents: 7399
diff changeset
    21
     (*Using the locale constant "f", this is  f (h (x,y))) = x*)
7badd511844d working snapshot
paulson
parents: 7399
diff changeset
    22
  
6297
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    23
  extend_set :: "['a*'b => 'c, 'a set] => 'c set"
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10064
diff changeset
    24
    "extend_set h A == h ` (A <*> UNIV)"
6297
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    25
7342
532841541d73 project constants
paulson
parents: 6677
diff changeset
    26
  project_set :: "['a*'b => 'c, 'c set] => 'a set"
532841541d73 project constants
paulson
parents: 6677
diff changeset
    27
    "project_set h C == {x. EX y. h(x,y) : C}"
532841541d73 project constants
paulson
parents: 6677
diff changeset
    28
532841541d73 project constants
paulson
parents: 6677
diff changeset
    29
  extend_act :: "['a*'b => 'c, ('a*'a) set] => ('c*'c) set"
7826
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7546
diff changeset
    30
    "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
    31
7878
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7826
diff changeset
    32
  project_act :: "['a*'b => 'c, ('c*'c) set] => ('a*'a) set"
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7826
diff changeset
    33
    "project_act h act == {(x,x'). EX y y'. (h(x,y), h(x',y')) : act}"
7342
532841541d73 project constants
paulson
parents: 6677
diff changeset
    34
6297
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    35
  extend :: "['a*'b => 'c, 'a program] => 'c program"
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    36
    "extend h F == mk_program (extend_set h (Init F),
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10064
diff changeset
    37
			       extend_act h ` Acts F,
a7897aebbffc *** empty log message ***
nipkow
parents: 10064
diff changeset
    38
			       project_act h -` AllowedActs F)"
6297
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    39
7878
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7826
diff changeset
    40
  (*Argument C allows weak safety laws to be projected*)
7880
62fb24e28e5e exchanged the first two args of "project" and "drop_prog"
paulson
parents: 7878
diff changeset
    41
  project :: "['a*'b => 'c, 'c set, 'c program] => 'a program"
10064
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 8948
diff changeset
    42
    "project h C F ==
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 8948
diff changeset
    43
       mk_program (project_set h (Init F),
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10064
diff changeset
    44
		   project_act h ` Restrict C ` Acts F,
10064
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 8948
diff changeset
    45
		   {act. Restrict (project_set h C) act :
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10064
diff changeset
    46
		         project_act h ` Restrict C ` AllowedActs F})"
7342
532841541d73 project constants
paulson
parents: 6677
diff changeset
    47
6297
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    48
locale Extend =
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    49
  fixes 
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    50
    f       :: 'c => 'a
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    51
    g       :: 'c => 'b
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    52
    h       :: "'a*'b => 'c"    (*isomorphism between 'a * 'b and 'c *)
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    53
    slice   :: ['c set, 'b] => 'a set
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    54
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    55
  assumes
7482
7badd511844d working snapshot
paulson
parents: 7399
diff changeset
    56
    good_h  "good_map h"
6297
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    57
  defines
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    58
    f_def       "f z == fst (inv h z)"
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    59
    g_def       "g z == snd (inv h z)"
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    60
    slice_def   "slice Z y == {x. h(x,y) : Z}"
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    61
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    62
end