src/HOL/Lambda/Commutation.thy
author paulson
Thu, 01 Oct 1998 18:30:05 +0200
changeset 5601 b6456ccd9e3e
parent 3439 54785105178c
child 8624 69619f870939
permissions -rw-r--r--
revised for new treatment of integers
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1476
608483c2122a expanded tabs; incorporated Konrad's changes
clasohm
parents: 1302
diff changeset
     1
(*  Title:      HOL/Lambda/Commutation.thy
1278
7e6189fc833c Commutation replaces Confluence
nipkow
parents:
diff changeset
     2
    ID:         $Id$
1476
608483c2122a expanded tabs; incorporated Konrad's changes
clasohm
parents: 1302
diff changeset
     3
    Author:     Tobias Nipkow
1278
7e6189fc833c Commutation replaces Confluence
nipkow
parents:
diff changeset
     4
    Copyright   1995  TU Muenchen
7e6189fc833c Commutation replaces Confluence
nipkow
parents:
diff changeset
     5
7e6189fc833c Commutation replaces Confluence
nipkow
parents:
diff changeset
     6
Abstract commutation and confluence notions.
7e6189fc833c Commutation replaces Confluence
nipkow
parents:
diff changeset
     7
*)
7e6189fc833c Commutation replaces Confluence
nipkow
parents:
diff changeset
     8
7e6189fc833c Commutation replaces Confluence
nipkow
parents:
diff changeset
     9
Commutation = Trancl +
7e6189fc833c Commutation replaces Confluence
nipkow
parents:
diff changeset
    10
7e6189fc833c Commutation replaces Confluence
nipkow
parents:
diff changeset
    11
consts
7e6189fc833c Commutation replaces Confluence
nipkow
parents:
diff changeset
    12
  square  :: "[('a*'a)set,('a*'a)set,('a*'a)set,('a*'a)set] => bool"
7e6189fc833c Commutation replaces Confluence
nipkow
parents:
diff changeset
    13
  commute :: "[('a*'a)set,('a*'a)set] => bool"
7e6189fc833c Commutation replaces Confluence
nipkow
parents:
diff changeset
    14
  confluent, diamond, Church_Rosser :: "('a*'a)set => bool"
7e6189fc833c Commutation replaces Confluence
nipkow
parents:
diff changeset
    15
7e6189fc833c Commutation replaces Confluence
nipkow
parents:
diff changeset
    16
defs
7e6189fc833c Commutation replaces Confluence
nipkow
parents:
diff changeset
    17
  square_def
7e6189fc833c Commutation replaces Confluence
nipkow
parents:
diff changeset
    18
 "square R S T U == !x y.(x,y):R --> (!z.(x,z):S --> (? u. (y,u):T & (z,u):U))"
7e6189fc833c Commutation replaces Confluence
nipkow
parents:
diff changeset
    19
7e6189fc833c Commutation replaces Confluence
nipkow
parents:
diff changeset
    20
  commute_def "commute R S == square R S S R"
7e6189fc833c Commutation replaces Confluence
nipkow
parents:
diff changeset
    21
  diamond_def "diamond R   == commute R R"
7e6189fc833c Commutation replaces Confluence
nipkow
parents:
diff changeset
    22
7e6189fc833c Commutation replaces Confluence
nipkow
parents:
diff changeset
    23
  Church_Rosser_def "Church_Rosser(R) ==   
3439
54785105178c converse -> ^-1
nipkow
parents: 1476
diff changeset
    24
  !x y. (x,y) : (R Un R^-1)^* --> (? z. (x,z) : R^* & (y,z) : R^*)"
1278
7e6189fc833c Commutation replaces Confluence
nipkow
parents:
diff changeset
    25
7e6189fc833c Commutation replaces Confluence
nipkow
parents:
diff changeset
    26
translations
7e6189fc833c Commutation replaces Confluence
nipkow
parents:
diff changeset
    27
  "confluent R" == "diamond(R^*)"
7e6189fc833c Commutation replaces Confluence
nipkow
parents:
diff changeset
    28
7e6189fc833c Commutation replaces Confluence
nipkow
parents:
diff changeset
    29
end