src/HOL/UNITY/Union.thy
author paulson
Thu, 11 Nov 1999 10:25:17 +0100
changeset 8005 b64d86018785
parent 7947 b999c1ab9327
child 8055 bb15396278fb
permissions -rw-r--r--
new-style infix declaration for "image"
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5252
1b0f14d11142 Union primitives and examples
paulson
parents:
diff changeset
     1
(*  Title:      HOL/UNITY/Union.thy
1b0f14d11142 Union primitives and examples
paulson
parents:
diff changeset
     2
    ID:         $Id$
1b0f14d11142 Union primitives and examples
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
1b0f14d11142 Union primitives and examples
paulson
parents:
diff changeset
     4
    Copyright   1998  University of Cambridge
1b0f14d11142 Union primitives and examples
paulson
parents:
diff changeset
     5
1b0f14d11142 Union primitives and examples
paulson
parents:
diff changeset
     6
Unions of programs
1b0f14d11142 Union primitives and examples
paulson
parents:
diff changeset
     7
5804
8e0a4c4fd67b Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents: 5648
diff changeset
     8
Partly from Misra's Chapter 5: Asynchronous Compositions of Programs
7359
98a2afab3f86 extra syntax for JN, making it more like UN
paulson
parents: 6295
diff changeset
     9
98a2afab3f86 extra syntax for JN, making it more like UN
paulson
parents: 6295
diff changeset
    10
Do we need a Meet operator?  (Aka Intersection)
7947
b999c1ab9327 working again; new treatment of LocalTo
paulson
parents: 7915
diff changeset
    11
b999c1ab9327 working again; new treatment of LocalTo
paulson
parents: 7915
diff changeset
    12
CAN PROBABLY DELETE the "Disjoint" predicate
5252
1b0f14d11142 Union primitives and examples
paulson
parents:
diff changeset
    13
*)
1b0f14d11142 Union primitives and examples
paulson
parents:
diff changeset
    14
1b0f14d11142 Union primitives and examples
paulson
parents:
diff changeset
    15
Union = SubstAx + FP +
1b0f14d11142 Union primitives and examples
paulson
parents:
diff changeset
    16
1b0f14d11142 Union primitives and examples
paulson
parents:
diff changeset
    17
constdefs
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5611
diff changeset
    18
  JOIN  :: ['a set, 'a => 'b program] => 'b program
6295
351b3c2b0d83 removed the infernal States, eqStates, compatible, etc.
paulson
parents: 6012
diff changeset
    19
    "JOIN I F == mk_program (INT i:I. Init (F i), UN i:I. Acts (F i))"
5252
1b0f14d11142 Union primitives and examples
paulson
parents:
diff changeset
    20
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5611
diff changeset
    21
  Join :: ['a program, 'a program] => 'a program      (infixl 65)
6295
351b3c2b0d83 removed the infernal States, eqStates, compatible, etc.
paulson
parents: 6012
diff changeset
    22
    "F Join G == mk_program (Init F Int Init G, Acts F Un Acts G)"
5252
1b0f14d11142 Union primitives and examples
paulson
parents:
diff changeset
    23
6295
351b3c2b0d83 removed the infernal States, eqStates, compatible, etc.
paulson
parents: 6012
diff changeset
    24
  SKIP :: 'a program
351b3c2b0d83 removed the infernal States, eqStates, compatible, etc.
paulson
parents: 6012
diff changeset
    25
    "SKIP == mk_program (UNIV, {})"
5259
86d80749453f Null program and a few new results
paulson
parents: 5252
diff changeset
    26
7878
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7826
diff changeset
    27
  Diff :: "['a set, 'a program, ('a * 'a)set set] => 'a program"
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7826
diff changeset
    28
    "Diff C G acts ==
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7826
diff changeset
    29
       mk_program (Init G, (Restrict C `` Acts G) - (Restrict C `` acts))"
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5611
diff changeset
    30
5804
8e0a4c4fd67b Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents: 5648
diff changeset
    31
  (*The set of systems that regard "v" as local to F*)
7878
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7826
diff changeset
    32
  LOCALTO :: ['a => 'b, 'a set, 'a program] => 'a program set
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7826
diff changeset
    33
                                           ("(_/ localTo[_]/ _)" [80,0,80] 80)
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7826
diff changeset
    34
    "v localTo[C] F == {G. ALL z. Diff C G (Acts F) : stable {s. v s = z}}"
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7826
diff changeset
    35
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7826
diff changeset
    36
  (*The weak version of localTo, considering only G's reachable states*)
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7826
diff changeset
    37
  LocalTo :: ['a => 'b, 'a program] => 'a program set  (infixl 80)
7915
c7fd7eb3b0ef ALMOST working version: LocalTo results commented out
paulson
parents: 7878
diff changeset
    38
    "v LocalTo F == {G. G : v localTo[reachable (F Join G)] F}"
5804
8e0a4c4fd67b Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents: 5648
diff changeset
    39
7826
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7359
diff changeset
    40
  (*Two programs with disjoint actions, except for identity actions.
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7359
diff changeset
    41
    It's a weak property but still useful.*)
7878
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7826
diff changeset
    42
  Disjoint :: ['a set, 'a program, 'a program] => bool
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7826
diff changeset
    43
    "Disjoint C F G ==
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7826
diff changeset
    44
       (Restrict C `` (Acts F - {Id})) Int (Restrict C `` (Acts G - {Id}))
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7826
diff changeset
    45
       <= {}"
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5611
diff changeset
    46
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents: 5259
diff changeset
    47
syntax
7359
98a2afab3f86 extra syntax for JN, making it more like UN
paulson
parents: 6295
diff changeset
    48
  "@JOIN1"     :: [pttrns, 'b set] => 'b set         ("(3JN _./ _)" 10)
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents: 5259
diff changeset
    49
  "@JOIN"      :: [pttrn, 'a set, 'b set] => 'b set  ("(3JN _:_./ _)" 10)
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents: 5259
diff changeset
    50
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents: 5259
diff changeset
    51
translations
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents: 5259
diff changeset
    52
  "JN x:A. B"   == "JOIN A (%x. B)"
7359
98a2afab3f86 extra syntax for JN, making it more like UN
paulson
parents: 6295
diff changeset
    53
  "JN x y. B"   == "JN x. JN y. B"
98a2afab3f86 extra syntax for JN, making it more like UN
paulson
parents: 6295
diff changeset
    54
  "JN x. B"     == "JOIN UNIV (%x. B)"
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents: 5259
diff changeset
    55
5252
1b0f14d11142 Union primitives and examples
paulson
parents:
diff changeset
    56
end