src/HOL/UNITY/Common.thy
author wenzelm
Mon, 13 Mar 2000 13:21:39 +0100
changeset 8434 5e4bba59bfaa
parent 7537 875754b599df
child 10064 1a77667b21ef
permissions -rw-r--r--
use HOLogic.Not; export indexify_names;

(*  Title:      HOL/UNITY/Common
    ID:         $Id$
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
    Copyright   1998  University of Cambridge

Common Meeting Time example from Misra (1994)

The state is identified with the one variable in existence.

From Misra, "A Logic for Concurrent Programming" (1994), sections 5.1 and 13.1.
*)

Common = Union + 

consts
  ftime,gtime :: nat=>nat

rules
  fmono "m <= n ==> ftime m <= ftime n"
  gmono "m <= n ==> gtime m <= gtime n"

  fasc  "m <= ftime n"
  gasc  "m <= gtime n"

constdefs
  common :: nat set
    "common == {n. ftime n = n & gtime n = n}"

  maxfg :: nat => nat set
    "maxfg m == {t. t <= max (ftime m) (gtime m)}"

end