--- a/src/Pure/build-jars Tue Jul 12 17:53:06 2011 +0200
+++ b/src/Pure/build-jars Tue Jul 12 18:00:05 2011 +0200
@@ -57,6 +57,7 @@
library.scala
package.scala
term.scala
+ term_xml.scala
)
--- a/src/Pure/term.scala Tue Jul 12 17:53:06 2011 +0200
+++ b/src/Pure/term.scala Tue Jul 12 18:00:05 2011 +0200
@@ -1,7 +1,7 @@
/* Title: Pure/term.scala
Author: Makarius
-Lambda terms with XML data representation.
+Lambda terms, types, sorts.
Note: Isabelle/ML is the primary environment for logical operations.
*/
@@ -11,8 +11,6 @@
object Term
{
- /* basic type definitions */
-
type Indexname = (String, Int)
type Sort = List[String]
@@ -33,55 +31,3 @@
case class App(fun: Term, arg: Term) extends Term
}
-
-object Term_XML
-{
- import Term._
-
-
- /* XML data representation */
-
- object Encode
- {
- import XML.Encode._
-
- val sort: T[Sort] = list(string)
-
- def typ: T[Typ] =
- variant[Typ](List(
- { case Type(a, b) => (List(a), list(typ)(b)) },
- { case TFree(a, b) => (List(a), sort(b)) },
- { case TVar((a, b), c) => (List(a, int_atom(b)), sort(c)) }))
-
- def term: T[Term] =
- variant[Term](List(
- { case Const(a, b) => (List(a), typ(b)) },
- { case Free(a, b) => (List(a), typ(b)) },
- { case Var((a, b), c) => (List(a, int_atom(b)), typ(c)) },
- { case Bound(a) => (List(int_atom(a)), Nil) },
- { case Abs(a, b, c) => (List(a), pair(typ, term)(b, c)) },
- { case App(a, b) => (Nil, pair(term, term)(a, b)) }))
- }
-
- object Decode
- {
- import XML.Decode._
-
- val sort: T[Sort] = list(string)
-
- def typ: T[Typ] =
- variant[Typ](List(
- { case (List(a), b) => Type(a, list(typ)(b)) },
- { case (List(a), b) => TFree(a, sort(b)) },
- { case (List(a, b), c) => TVar((a, int_atom(b)), sort(c)) }))
-
- def term: T[Term] =
- variant[Term](List(
- { case (List(a), b) => Const(a, typ(b)) },
- { case (List(a), b) => Free(a, typ(b)) },
- { case (List(a, b), c) => Var((a, int_atom(b)), typ(c)) },
- { case (List(a), Nil) => Bound(int_atom(a)) },
- { case (List(a), b) => val (c, d) = pair(typ, term)(b); Abs(a, c, d) },
- { case (Nil, a) => val (b, c) = pair(term, term)(a); App(b, c) }))
- }
-}
--- a/src/Pure/term_xml.ML Tue Jul 12 17:53:06 2011 +0200
+++ b/src/Pure/term_xml.ML Tue Jul 12 18:00:05 2011 +0200
@@ -14,15 +14,13 @@
signature TERM_XML =
sig
- structure Encode: TERM_XML_OPS where type 'a T = 'a XML.Encode.T
- structure Decode: TERM_XML_OPS where type 'a T = 'a XML.Decode.T
+ structure Encode: TERM_XML_OPS
+ structure Decode: TERM_XML_OPS
end;
structure Term_XML: TERM_XML =
struct
-(* encode *)
-
structure Encode =
struct
@@ -45,9 +43,6 @@
end;
-
-(* decode *)
-
structure Decode =
struct
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/term_xml.scala Tue Jul 12 18:00:05 2011 +0200
@@ -0,0 +1,57 @@
+/* Title: Pure/term_xml.scala
+ Author: Makarius
+
+XML data representation of lambda terms.
+*/
+
+package isabelle
+
+
+object Term_XML
+{
+ import Term._
+
+ object Encode
+ {
+ import XML.Encode._
+
+ val sort: T[Sort] = list(string)
+
+ def typ: T[Typ] =
+ variant[Typ](List(
+ { case Type(a, b) => (List(a), list(typ)(b)) },
+ { case TFree(a, b) => (List(a), sort(b)) },
+ { case TVar((a, b), c) => (List(a, int_atom(b)), sort(c)) }))
+
+ def term: T[Term] =
+ variant[Term](List(
+ { case Const(a, b) => (List(a), typ(b)) },
+ { case Free(a, b) => (List(a), typ(b)) },
+ { case Var((a, b), c) => (List(a, int_atom(b)), typ(c)) },
+ { case Bound(a) => (List(int_atom(a)), Nil) },
+ { case Abs(a, b, c) => (List(a), pair(typ, term)(b, c)) },
+ { case App(a, b) => (Nil, pair(term, term)(a, b)) }))
+ }
+
+ object Decode
+ {
+ import XML.Decode._
+
+ val sort: T[Sort] = list(string)
+
+ def typ: T[Typ] =
+ variant[Typ](List(
+ { case (List(a), b) => Type(a, list(typ)(b)) },
+ { case (List(a), b) => TFree(a, sort(b)) },
+ { case (List(a, b), c) => TVar((a, int_atom(b)), sort(c)) }))
+
+ def term: T[Term] =
+ variant[Term](List(
+ { case (List(a), b) => Const(a, typ(b)) },
+ { case (List(a), b) => Free(a, typ(b)) },
+ { case (List(a, b), c) => Var((a, int_atom(b)), typ(c)) },
+ { case (List(a), Nil) => Bound(int_atom(a)) },
+ { case (List(a), b) => val (c, d) = pair(typ, term)(b); Abs(a, c, d) },
+ { case (Nil, a) => val (b, c) = pair(term, term)(a); App(b, c) }))
+ }
+}