src/Pure/term.scala
author wenzelm
Sun, 10 Jul 2011 17:58:11 +0200
changeset 43730 a0ed7bc688b5
child 43731 70072780e095
permissions -rw-r--r--
lambda terms with XML data representation in Scala; avoid `class` in signature;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
43730
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
     1
/*  Title:      Pure/term.scala
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
     2
    Author:     Makarius
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
     3
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
     4
Lambda terms with XML data representation.
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
     5
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
     6
Note: Isabelle/ML is the primary environment for logical operations.
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
     7
*/
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
     8
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
     9
package isabelle
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    10
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    11
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    12
object Term
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    13
{
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    14
  /* basic type definitions */
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    15
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    16
  type Indexname = (String, Int)
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    17
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    18
  type Sort = List[String]
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    19
  val dummyS: Sort = List("")
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    20
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    21
  sealed abstract class Typ
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    22
  case class Type(name: String, args: List[Typ] = Nil) extends Typ
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    23
  case class TFree(name: String, sort: Sort = dummyS) extends Typ
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    24
  case class TVar(name: Indexname, sort: Sort = dummyS) extends Typ
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    25
  val dummyT = Type("dummy")
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    26
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    27
  sealed abstract class Term
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    28
  case class Const(name: String, typ: Typ = dummyT) extends Term
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    29
  case class Free(name: String, typ: Typ = dummyT) extends Term
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    30
  case class Var(name: Indexname, typ: Typ = dummyT) extends Term
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    31
  case class Bound(index: Int) extends Term
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    32
  case class Abs(name: String, typ: Typ = dummyT, body: Term) extends Term
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    33
  case class App(fun: Term, arg: Term) extends Term
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    34
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    35
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    36
  /* XML data representation */
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    37
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    38
  object XML
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    39
  {
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    40
    object Make
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    41
    {
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    42
      import XML_Data.Make._
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    43
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    44
      val indexname: T[Indexname] = pair(string, int)
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    45
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    46
      val sort: T[Sort] = list(string)
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    47
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    48
      def typ: T[Typ] =
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    49
        variant[Typ](List(
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    50
          { case Type(a, b) => pair(string, list(typ))((a, b)) },
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    51
          { case TFree(a, b) => pair(string, sort)((a, b)) },
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    52
          { case TVar(a, b) => pair(indexname, sort)((a, b)) }))
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    53
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    54
      def term: T[Term] =
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    55
        variant[Term](List(
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    56
          { case Const(a, b) => pair(string, typ)((a, b)) },
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    57
          { case Free(a, b) => pair(string, typ)((a, b)) },
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    58
          { case Var(a, b) => pair(indexname, typ)((a, b)) },
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    59
          { case Bound(a) => int(a) },
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    60
          { case Abs(a, b, c) => triple(string, typ, term)((a, b, c)) },
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    61
          { case App(a, b) => pair(term, term)((a, b)) }))
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    62
    }
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    63
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    64
    object Dest
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    65
    {
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    66
      import XML_Data.Dest._
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    67
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    68
      val indexname: T[Indexname] = pair(string, int)
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    69
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    70
      val sort: T[Sort] = list(string)
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    71
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    72
      def typ: T[Typ] =
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    73
        variant[Typ](List(
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    74
          (x => { val (a, b) = pair(string, list(typ))(x); Type(a, b) }),
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    75
          (x => { val (a, b) = pair(string, sort)(x); TFree(a, b) }),
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    76
          (x => { val (a, b) = pair(indexname, sort)(x); TVar(a, b) })))
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    77
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    78
      def term: T[Term] =
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    79
        variant[Term](List(
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    80
          (x => { val (a, b) = pair(string, typ)(x); Const(a, b) }),
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    81
          (x => { val (a, b) = pair(string, typ)(x); Free(a, b) }),
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    82
          (x => { val (a, b) = pair(indexname, typ)(x); Var(a, b) }),
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    83
          (x => Bound(int(x))),
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    84
          (x => { val (a, b, c) = triple(string, typ, term)(x); Abs(a, b, c) }),
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    85
          (x => { val (a, b) = pair(term, term)(x); App(a, b) })))
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    86
    }
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    87
  }
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    88
}