src/HOL/UNITY/Lift_prog.thy
author paulson
Mon, 18 Oct 1999 16:16:03 +0200
changeset 7880 62fb24e28e5e
parent 7878 43b03d412b82
child 7947 b999c1ab9327
permissions -rw-r--r--
exchanged the first two args of "project" and "drop_prog"
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
7186
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
     1
(*  Title:      HOL/UNITY/Lift_prog.thy
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
     2
    ID:         $Id$
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
     4
    Copyright   1999  University of Cambridge
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
     5
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
     6
lift_prog, etc: replication of components
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
     7
*)
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
     8
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents: 7482
diff changeset
     9
Lift_prog = Project +
7186
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
    10
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
    11
constdefs
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
    12
7482
7badd511844d working snapshot
paulson
parents: 7399
diff changeset
    13
  lift_map :: "['a, 'b * ('a => 'b)] => ('a => 'b)"
7badd511844d working snapshot
paulson
parents: 7399
diff changeset
    14
    "lift_map i == %(s,f). f(i := s)"
7badd511844d working snapshot
paulson
parents: 7399
diff changeset
    15
7186
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
    16
  lift_set :: "['a, 'b set] => ('a => 'b) set"
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
    17
    "lift_set i A == {f. f i : A}"
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
    18
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
    19
  drop_set :: "['a, ('a=>'b) set] => 'b set"
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
    20
    "drop_set i A == (%f. f i) `` A"
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
    21
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
    22
  lift_act :: "['a, ('b*'b) set] => (('a=>'b) * ('a=>'b)) set"
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
    23
    "lift_act i act == {(f,f'). f(i:= f' i) = f' & (f i, f' i) : act}"
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
    24
7878
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7688
diff changeset
    25
  drop_act :: "['a, (('a=>'b) * ('a=>'b)) set] => ('b*'b) set"
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7688
diff changeset
    26
    "drop_act i act == {(f i, f' i) | f f'. (f,f'): act}"
7186
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
    27
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
    28
  lift_prog :: "['a, 'b program] => ('a => 'b) program"
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
    29
    "lift_prog i F ==
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
    30
       mk_program (lift_set i (Init F),
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
    31
		   lift_act i `` Acts F)"
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
    32
7878
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7688
diff changeset
    33
  (*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
    34
  drop_prog :: "['a, ('a=>'b) set, ('a=>'b) program] => 'b program"
62fb24e28e5e exchanged the first two args of "project" and "drop_prog"
paulson
parents: 7878
diff changeset
    35
    "drop_prog i C F ==
7186
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
    36
       mk_program (drop_set i (Init F),
7878
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7688
diff changeset
    37
		   drop_act i `` Restrict C `` (Acts F))"
7186
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
    38
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
    39
  (*simplifies the expression of specifications*)
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
    40
  constdefs
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
    41
    sub :: ['a, 'a=>'b] => 'b
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
    42
      "sub i f == f i"
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
    43
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
    44
860479291bb5 new theory UNITY/Lift_prog
paulson
parents:
diff changeset
    45
end