| author | paulson | 
| Wed, 19 Jun 2002 12:48:55 +0200 | |
| changeset 13225 | b6fc6e4a0a24 | 
| parent 12114 | a8e860c86252 | 
| child 13792 | d1811693899c | 
| permissions | -rw-r--r-- | 
| 5252 | 1 | (* Title: HOL/UNITY/Union.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 4 | Copyright 1998 University of Cambridge | |
| 5 | ||
| 6 | Unions of programs | |
| 7 | ||
| 5804 
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier.  New lemmas
 paulson parents: 
5648diff
changeset | 8 | Partly from Misra's Chapter 5: Asynchronous Compositions of Programs | 
| 5252 | 9 | *) | 
| 10 | ||
| 11 | Union = SubstAx + FP + | |
| 12 | ||
| 13 | constdefs | |
| 10064 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9685diff
changeset | 14 | |
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9685diff
changeset | 15 |   (*FIXME: conjoin Init F Int Init G ~= {} *) 
 | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9685diff
changeset | 16 | ok :: ['a program, 'a program] => bool (infixl 65) | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9685diff
changeset | 17 | "F ok G == Acts F <= AllowedActs G & | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9685diff
changeset | 18 | Acts G <= AllowedActs F" | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9685diff
changeset | 19 | |
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9685diff
changeset | 20 |   (*FIXME: conjoin (INT i:I. Init (F i)) ~= {} *) 
 | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9685diff
changeset | 21 | OK :: ['a set, 'a => 'b program] => bool | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9685diff
changeset | 22 |     "OK I F == (ALL i:I. ALL j: I-{i}. Acts (F i) <= AllowedActs (F j))"
 | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9685diff
changeset | 23 | |
| 5648 | 24 | JOIN :: ['a set, 'a => 'b program] => 'b program | 
| 10064 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9685diff
changeset | 25 | "JOIN I F == mk_program (INT i:I. Init (F i), UN i:I. Acts (F i), | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9685diff
changeset | 26 | INT i:I. AllowedActs (F i))" | 
| 5252 | 27 | |
| 5648 | 28 | Join :: ['a program, 'a program] => 'a program (infixl 65) | 
| 10064 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9685diff
changeset | 29 | "F Join G == mk_program (Init F Int Init G, Acts F Un Acts G, | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9685diff
changeset | 30 | AllowedActs F Int AllowedActs G)" | 
| 5252 | 31 | |
| 6295 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6012diff
changeset | 32 | SKIP :: 'a program | 
| 10064 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9685diff
changeset | 33 |     "SKIP == mk_program (UNIV, {}, UNIV)"
 | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9685diff
changeset | 34 | |
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9685diff
changeset | 35 | (*Characterizes safety properties. Used with specifying AllowedActs*) | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9685diff
changeset | 36 | safety_prop :: "'a program set => bool" | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9685diff
changeset | 37 | "safety_prop X == SKIP: X & (ALL G. Acts G <= UNION X Acts --> G : X)" | 
| 5259 | 38 | |
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5259diff
changeset | 39 | syntax | 
| 7359 | 40 |   "@JOIN1"     :: [pttrns, 'b set] => 'b set         ("(3JN _./ _)" 10)
 | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5259diff
changeset | 41 |   "@JOIN"      :: [pttrn, 'a set, 'b set] => 'b set  ("(3JN _:_./ _)" 10)
 | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5259diff
changeset | 42 | |
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5259diff
changeset | 43 | translations | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5259diff
changeset | 44 | "JN x:A. B" == "JOIN A (%x. B)" | 
| 7359 | 45 | "JN x y. B" == "JN x. JN y. B" | 
| 46 | "JN x. B" == "JOIN UNIV (%x. B)" | |
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5259diff
changeset | 47 | |
| 12114 
a8e860c86252
eliminated old "symbols" syntax, use "xsymbols" instead;
 wenzelm parents: 
10064diff
changeset | 48 | syntax (xsymbols) | 
| 9685 | 49 |   SKIP      :: 'a program                              ("\\<bottom>")
 | 
| 50 | "op Join" :: ['a program, 'a program] => 'a program (infixl "\\<squnion>" 65) | |
| 51 |   "@JOIN1"  :: [pttrns, 'b set] => 'b set              ("(3\\<Squnion> _./ _)" 10)
 | |
| 52 |   "@JOIN"   :: [pttrn, 'a set, 'b set] => 'b set       ("(3\\<Squnion> _:_./ _)" 10)
 | |
| 53 | ||
| 5252 | 54 | end |