src/HOL/Subst/Subst.thy
author berghofe
Fri, 24 Jul 1998 13:19:38 +0200
changeset 5184 9b8547a9496a
parent 3842 b55686a7b22c
child 15635 8408a06590a6
permissions -rw-r--r--
Adapted to new datatype package.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
3268
012c43174664 Mostly cosmetic changes: updated headers, ID lines, etc.
paulson
parents: 3192
diff changeset
     1
(*  Title:      Subst/Subst.thy
012c43174664 Mostly cosmetic changes: updated headers, ID lines, etc.
paulson
parents: 3192
diff changeset
     2
    ID:         $Id$
1476
608483c2122a expanded tabs; incorporated Konrad's changes
clasohm
parents: 1266
diff changeset
     3
    Author:     Martin Coen, Cambridge University Computer Laboratory
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
     4
    Copyright   1993  University of Cambridge
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
     5
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
     6
Substitutions on uterms
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
     7
*)
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
     8
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 1476
diff changeset
     9
Subst = AList + UTerm +
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    10
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    11
consts
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    12
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 1476
diff changeset
    13
  "=$="  ::  "[('a*('a uterm)) list,('a*('a uterm)) list] => bool" (infixr 52)
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 1476
diff changeset
    14
  "<|"   ::  "'a uterm => ('a * 'a uterm) list => 'a uterm"        (infixl 55)
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 1476
diff changeset
    15
  "<>"   ::  "[('a*('a uterm)) list, ('a*('a uterm)) list] 
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 1476
diff changeset
    16
                 => ('a*('a uterm)) list"                          (infixl 56)
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    17
  sdom   ::  "('a*('a uterm)) list => 'a set"
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    18
  srange ::  "('a*('a uterm)) list => 'a set"
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    19
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    20
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 3842
diff changeset
    21
primrec
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 1476
diff changeset
    22
  subst_Var      "(Var v <| s) = assoc v (Var v) s"
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 1476
diff changeset
    23
  subst_Const  "(Const c <| s) = Const c"
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 1476
diff changeset
    24
  subst_Comb  "(Comb M N <| s) = Comb (M <| s) (N <| s)"
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    25
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 1476
diff changeset
    26
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 1476
diff changeset
    27
defs 
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    28
3842
b55686a7b22c fixed dots;
wenzelm
parents: 3268
diff changeset
    29
  subst_eq_def  "r =$= s == ALL t. t <| r = t <| s"
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 1476
diff changeset
    30
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 1476
diff changeset
    31
  comp_def    "al <> bl == alist_rec al bl (%x y xs g. (x,y <| bl)#g)"
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
  sdom_def
1151
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 972
diff changeset
    34
  "sdom(al) == alist_rec al {}  
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 1476
diff changeset
    35
                (%x y xs g. if Var(x)=y then g - {x} else g Un {x})"
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 1476
diff changeset
    36
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    37
  srange_def   
3192
a75558a4ed37 New version, modified by Konrad Slind and LCP for TFL
paulson
parents: 1476
diff changeset
    38
   "srange(al) == Union({y. EX x:sdom(al). y=vars_of(Var(x) <| al)})"
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    39
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    40
end