src/HOL/UNITY/Comp.thy
author wenzelm
Fri, 09 Nov 2001 00:09:47 +0100
changeset 12114 a8e860c86252
parent 11190 44e157622cb2
child 12338 de0f4a63baa5
permissions -rw-r--r--
eliminated old "symbols" syntax, use "xsymbols" instead;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5597
a12b25c53df1 composition theory
paulson
parents:
diff changeset
     1
(*  Title:      HOL/UNITY/Comp.thy
a12b25c53df1 composition theory
paulson
parents:
diff changeset
     2
    ID:         $Id$
a12b25c53df1 composition theory
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
a12b25c53df1 composition theory
paulson
parents:
diff changeset
     4
    Copyright   1998  University of Cambridge
a12b25c53df1 composition theory
paulson
parents:
diff changeset
     5
a12b25c53df1 composition theory
paulson
parents:
diff changeset
     6
Composition
11190
44e157622cb2 *** empty log message ***
ehmety
parents: 8128
diff changeset
     7
From Chandy and Sanders, "Reasoning About Program Composition",
44e157622cb2 *** empty log message ***
ehmety
parents: 8128
diff changeset
     8
Technical Report 2000-003, University of Florida, 2000.
5597
a12b25c53df1 composition theory
paulson
parents:
diff changeset
     9
11190
44e157622cb2 *** empty log message ***
ehmety
parents: 8128
diff changeset
    10
Revised by Sidi Ehmety on January  2001 
44e157622cb2 *** empty log message ***
ehmety
parents: 8128
diff changeset
    11
44e157622cb2 *** empty log message ***
ehmety
parents: 8128
diff changeset
    12
Added: a strong form of the <= relation (component_of) and localize 
44e157622cb2 *** empty log message ***
ehmety
parents: 8128
diff changeset
    13
5597
a12b25c53df1 composition theory
paulson
parents:
diff changeset
    14
*)
a12b25c53df1 composition theory
paulson
parents:
diff changeset
    15
a12b25c53df1 composition theory
paulson
parents:
diff changeset
    16
Comp = Union +
a12b25c53df1 composition theory
paulson
parents:
diff changeset
    17
7399
cf780c2bcccf changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
paulson
parents: 7364
diff changeset
    18
instance
cf780c2bcccf changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
paulson
parents: 7364
diff changeset
    19
  program :: (term)ord
5597
a12b25c53df1 composition theory
paulson
parents:
diff changeset
    20
7399
cf780c2bcccf changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
paulson
parents: 7364
diff changeset
    21
defs
cf780c2bcccf changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
paulson
parents: 7364
diff changeset
    22
  component_def   "F <= H == EX G. F Join G = H"
cf780c2bcccf changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
paulson
parents: 7364
diff changeset
    23
  strict_component_def   "(F < (H::'a program)) == (F <= H & F ~= H)"
5612
e981ca6f7332 Finished proofs to end of section 5.1 of Chandy and Sanders
paulson
parents: 5597
diff changeset
    24
11190
44e157622cb2 *** empty log message ***
ehmety
parents: 8128
diff changeset
    25
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7399
diff changeset
    26
constdefs
11190
44e157622cb2 *** empty log message ***
ehmety
parents: 8128
diff changeset
    27
  component_of :: "'a program=>'a program=> bool"
44e157622cb2 *** empty log message ***
ehmety
parents: 8128
diff changeset
    28
                                    (infixl "component'_of" 50)
44e157622cb2 *** empty log message ***
ehmety
parents: 8128
diff changeset
    29
  "F component_of H == EX G. F ok G & F Join G = H"
44e157622cb2 *** empty log message ***
ehmety
parents: 8128
diff changeset
    30
44e157622cb2 *** empty log message ***
ehmety
parents: 8128
diff changeset
    31
  strict_component_of :: "'a program\\<Rightarrow>'a program=> bool"
44e157622cb2 *** empty log message ***
ehmety
parents: 8128
diff changeset
    32
                                    (infixl "strict'_component'_of" 50)
44e157622cb2 *** empty log message ***
ehmety
parents: 8128
diff changeset
    33
  "F strict_component_of H == F component_of H & F~=H"
44e157622cb2 *** empty log message ***
ehmety
parents: 8128
diff changeset
    34
  
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7399
diff changeset
    35
  preserves :: "('a=>'b) => 'a program set"
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7399
diff changeset
    36
    "preserves v == INT z. stable {s. v s = z}"
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7399
diff changeset
    37
11190
44e157622cb2 *** empty log message ***
ehmety
parents: 8128
diff changeset
    38
  localize  :: "('a=>'b) => 'a program => 'a program"
44e157622cb2 *** empty log message ***
ehmety
parents: 8128
diff changeset
    39
  "localize v F == mk_program(Init F, Acts F,
44e157622cb2 *** empty log message ***
ehmety
parents: 8128
diff changeset
    40
			      AllowedActs F Int (UN G:preserves v. Acts G))"
44e157622cb2 *** empty log message ***
ehmety
parents: 8128
diff changeset
    41
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7399
diff changeset
    42
  funPair      :: "['a => 'b, 'a => 'c, 'a] => 'b * 'c"
11190
44e157622cb2 *** empty log message ***
ehmety
parents: 8128
diff changeset
    43
  "funPair f g == %x. (f x, g x)"
5597
a12b25c53df1 composition theory
paulson
parents:
diff changeset
    44
end