src/ZF/UNITY/Constrains.thy
author paulson
Tue, 27 May 2003 11:39:03 +0200
changeset 14046 6616e6c53d48
parent 12195 ed2893765a08
child 15634 bca33c49b083
permissions -rw-r--r--
updating ZF-UNITY with Sidi's new material
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     1
(*  Title:      ZF/UNITY/Constrains.thy
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     2
    ID:         $Id$
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     3
    Author:     Sidi O Ehmety, Computer Laboratory
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     4
    Copyright   2001  University of Cambridge
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     5
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     6
Safety relations: restricted to the set of reachable states.
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     7
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     8
Theory ported from HOL.
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     9
*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    10
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    11
Constrains = UNITY +
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    12
consts traces :: "[i, i] => i"
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    13
  (* Initial states and program => (final state, reversed trace to it)... 
12195
ed2893765a08 *** empty log message ***
ehmety
parents: 11479
diff changeset
    14
      the domain may also be state*list(state) *)
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    15
inductive 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    16
  domains 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    17
     "traces(init, acts)" <=
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    18
         "(init Un (UN act:acts. field(act)))*list(UN act:acts. field(act))"
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    19
  intrs 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    20
         (*Initial trace is empty*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    21
    Init  "s: init ==> <s,[]> : traces(init,acts)"
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    22
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    23
    Acts  "[| act:acts;  <s,evs> : traces(init,acts);  <s,s'>: act |]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    24
           ==> <s', Cons(s,evs)> : traces(init, acts)"
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    25
  
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    26
  type_intrs "list.intrs@[UnI1, UnI2, UN_I, fieldI2, fieldI1]"
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    27
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    28
  consts reachable :: "i=>i"
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    29
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    30
inductive
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    31
  domains
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    32
  "reachable(F)" <= "Init(F) Un (UN act:Acts(F). field(act))"
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    33
  intrs 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    34
    Init  "s:Init(F) ==> s:reachable(F)"
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    35
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    36
    Acts  "[| act: Acts(F);  s:reachable(F);  <s,s'>: act |]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    37
           ==> s':reachable(F)"
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    38
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    39
  type_intrs "[UnI1, UnI2, fieldI2, UN_I]"
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    40
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    41
  
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    42
consts
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    43
  Constrains :: "[i,i] => i"  (infixl "Co"     60)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    44
  op_Unless  :: "[i, i] => i"  (infixl "Unless" 60)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    45
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    46
defs
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    47
  Constrains_def
12195
ed2893765a08 *** empty log message ***
ehmety
parents: 11479
diff changeset
    48
    "A Co B == {F:program. F:(reachable(F) Int A) co B}"
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    49
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    50
  Unless_def
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    51
    "A Unless B == (A-B) Co (A Un B)"
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    52
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    53
constdefs
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    54
  Stable     :: "i => i"
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    55
    "Stable(A) == A Co A"
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    56
  (*Always is the weak form of "invariant"*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    57
  Always :: "i => i"
12195
ed2893765a08 *** empty log message ***
ehmety
parents: 11479
diff changeset
    58
    "Always(A) == initially(A) Int Stable(A)"
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    59
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    60
end
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    61