src/HOL/Subst/Unifier.thy
author clasohm
Mon, 05 Feb 1996 21:29:06 +0100
changeset 1476 608483c2122a
parent 1374 5e407f2a3323
child 3192 a75558a4ed37
permissions -rw-r--r--
expanded tabs; incorporated Konrad's changes
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1476
608483c2122a expanded tabs; incorporated Konrad's changes
clasohm
parents: 1374
diff changeset
     1
(*  Title:      Subst/unifier.thy
608483c2122a expanded tabs; incorporated Konrad's changes
clasohm
parents: 1374
diff changeset
     2
    Author:     Martin Coen, Cambridge University Computer Laboratory
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
     3
    Copyright   1993  University of Cambridge
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
     4
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
     5
Definition of most general idempotent unifier
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
     6
*)
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
     7
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
     8
Unifier = Subst +
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
     9
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    10
consts
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    11
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    12
  Idem       :: "('a*('a uterm))list=> bool"
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    13
  Unifier    :: "[('a*('a uterm))list,'a uterm,'a uterm] => bool"
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    14
  ">>"       :: "[('a*('a uterm))list,('a*('a uterm))list] => bool" (infixr 52)
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    15
  MGUnifier  :: "[('a*('a uterm))list,'a uterm,'a uterm] => bool"
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    16
  MGIUnifier :: "[('a*('a uterm))list,'a uterm,'a uterm] => bool"
1374
5e407f2a3323 removed quotes from consts and syntax sections
clasohm
parents: 1151
diff changeset
    17
  UWFD       :: ['a uterm,'a uterm,'a uterm,'a uterm] => bool
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    18
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    19
rules  (*Definitions*)
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    20
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    21
  Idem_def         "Idem(s) == s <> s =s= s"
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    22
  Unifier_def      "Unifier s t u == t <| s = u <| s"
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    23
  MoreGeneral_def  "r >> s == ? q.s =s= r <> q"
1151
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 968
diff changeset
    24
  MGUnifier_def    "MGUnifier s t u == Unifier s t u & 
1476
608483c2122a expanded tabs; incorporated Konrad's changes
clasohm
parents: 1374
diff changeset
    25
                    (! r.Unifier r t u --> s >> r)"
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    26
  MGIUnifier_def   "MGIUnifier s t u == MGUnifier s t u & Idem(s)"
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    27
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    28
  UWFD_def
1151
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 968
diff changeset
    29
  "UWFD x y x' y' == 
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 968
diff changeset
    30
    (vars_of(x) Un vars_of(y) < vars_of(x') Un vars_of(y')) |  
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 968
diff changeset
    31
    (vars_of(x) Un vars_of(y) = vars_of(x') Un vars_of(y') & x <: x')"
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    32
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    33
end