src/HOL/Subst/Unify.thy
author paulson
Wed, 05 Nov 1997 13:23:46 +0100
changeset 4153 e534c4c32d54
parent 3298 5f0ed3caa991
child 6481 dbf2d9b3d6c8
permissions -rw-r--r--
Ran expandshort, especially to introduce Safe_tac
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents:
diff changeset
     1
(*  Title:      Subst/Unify
3267
7203f4dbc0c5 Removed rprod from the WF relation
paulson
parents: 3192
diff changeset
     2
    ID:         $Id$
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents:
diff changeset
     3
    Author:     Konrad Slind, Cambridge University Computer Laboratory
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents:
diff changeset
     4
    Copyright   1997  University of Cambridge
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents:
diff changeset
     5
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents:
diff changeset
     6
Unification algorithm
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents:
diff changeset
     7
*)
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents:
diff changeset
     8
3298
5f0ed3caa991 Now uses type option instead of Fail/Subst
paulson
parents: 3267
diff changeset
     9
Unify = Unifier + WF_Rel + Option + 
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents:
diff changeset
    10
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents:
diff changeset
    11
consts
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents:
diff changeset
    12
3267
7203f4dbc0c5 Removed rprod from the WF relation
paulson
parents: 3192
diff changeset
    13
   unifyRel :: "(('a uterm * 'a uterm) * ('a uterm * 'a uterm)) set"
3298
5f0ed3caa991 Now uses type option instead of Fail/Subst
paulson
parents: 3267
diff changeset
    14
   unify    :: "'a uterm * 'a uterm => ('a * 'a uterm) list option"
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents:
diff changeset
    15
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents:
diff changeset
    16
defs
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents:
diff changeset
    17
3267
7203f4dbc0c5 Removed rprod from the WF relation
paulson
parents: 3192
diff changeset
    18
  (*Termination relation for the Unify function:
7203f4dbc0c5 Removed rprod from the WF relation
paulson
parents: 3192
diff changeset
    19
    --either the set of variables decreases
7203f4dbc0c5 Removed rprod from the WF relation
paulson
parents: 3192
diff changeset
    20
    --or the first argument does (in fact, both do)
7203f4dbc0c5 Removed rprod from the WF relation
paulson
parents: 3192
diff changeset
    21
  *)
7203f4dbc0c5 Removed rprod from the WF relation
paulson
parents: 3192
diff changeset
    22
  unifyRel_def  "unifyRel == inv_image  (finite_psubset ** measure uterm_size)
7203f4dbc0c5 Removed rprod from the WF relation
paulson
parents: 3192
diff changeset
    23
                               (%(M,N). (vars_of M Un vars_of N, M))"
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents:
diff changeset
    24
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents:
diff changeset
    25
recdef unify "unifyRel"
3298
5f0ed3caa991 Now uses type option instead of Fail/Subst
paulson
parents: 3267
diff changeset
    26
  "unify(Const m, Const n)  = (if (m=n) then Some[] else None)"
5f0ed3caa991 Now uses type option instead of Fail/Subst
paulson
parents: 3267
diff changeset
    27
  "unify(Const m, Comb M N) = None"
5f0ed3caa991 Now uses type option instead of Fail/Subst
paulson
parents: 3267
diff changeset
    28
  "unify(Const m, Var v)    = Some[(v,Const m)]"
5f0ed3caa991 Now uses type option instead of Fail/Subst
paulson
parents: 3267
diff changeset
    29
  "unify(Var v, M)          = (if (Var v <: M) then None else Some[(v,M)])"
5f0ed3caa991 Now uses type option instead of Fail/Subst
paulson
parents: 3267
diff changeset
    30
  "unify(Comb M N, Const x) = None"
5f0ed3caa991 Now uses type option instead of Fail/Subst
paulson
parents: 3267
diff changeset
    31
  "unify(Comb M N, Var v)   = (if (Var v <: Comb M N) then None   
5f0ed3caa991 Now uses type option instead of Fail/Subst
paulson
parents: 3267
diff changeset
    32
                               else Some[(v,Comb M N)])"
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents:
diff changeset
    33
  "unify(Comb M1 N1, Comb M2 N2) =   
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents:
diff changeset
    34
      (case unify(M1,M2)  
3298
5f0ed3caa991 Now uses type option instead of Fail/Subst
paulson
parents: 3267
diff changeset
    35
        of None       => None  
5f0ed3caa991 Now uses type option instead of Fail/Subst
paulson
parents: 3267
diff changeset
    36
         | Some theta => (case unify(N1 <| theta, N2 <| theta)  
5f0ed3caa991 Now uses type option instead of Fail/Subst
paulson
parents: 3267
diff changeset
    37
                            of None       => None  
5f0ed3caa991 Now uses type option instead of Fail/Subst
paulson
parents: 3267
diff changeset
    38
                             | Some sigma => Some (theta <> sigma)))"
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents:
diff changeset
    39
end