src/Pure/term_xml.ML
author traytel
Mon, 24 Oct 2016 16:53:32 +0200
changeset 64379 71f42dcaa1df
parent 43789 321ebd051897
child 70784 799437173553
permissions -rw-r--r--
additional user-specified simp (naturality) rules used in friend_of_corec
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
43729
07d3c6afa865 XML data representation of lambda terms;
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/term_xml.ML
07d3c6afa865 XML data representation of lambda terms;
wenzelm
parents:
diff changeset
     2
    Author:     Makarius
07d3c6afa865 XML data representation of lambda terms;
wenzelm
parents:
diff changeset
     3
07d3c6afa865 XML data representation of lambda terms;
wenzelm
parents:
diff changeset
     4
XML data representation of lambda terms.
07d3c6afa865 XML data representation of lambda terms;
wenzelm
parents:
diff changeset
     5
*)
07d3c6afa865 XML data representation of lambda terms;
wenzelm
parents:
diff changeset
     6
07d3c6afa865 XML data representation of lambda terms;
wenzelm
parents:
diff changeset
     7
signature TERM_XML_OPS =
07d3c6afa865 XML data representation of lambda terms;
wenzelm
parents:
diff changeset
     8
sig
07d3c6afa865 XML data representation of lambda terms;
wenzelm
parents:
diff changeset
     9
  type 'a T
07d3c6afa865 XML data representation of lambda terms;
wenzelm
parents:
diff changeset
    10
  val sort: sort T
07d3c6afa865 XML data representation of lambda terms;
wenzelm
parents:
diff changeset
    11
  val typ: typ T
07d3c6afa865 XML data representation of lambda terms;
wenzelm
parents:
diff changeset
    12
  val term: term T
07d3c6afa865 XML data representation of lambda terms;
wenzelm
parents:
diff changeset
    13
end
07d3c6afa865 XML data representation of lambda terms;
wenzelm
parents:
diff changeset
    14
07d3c6afa865 XML data representation of lambda terms;
wenzelm
parents:
diff changeset
    15
signature TERM_XML =
07d3c6afa865 XML data representation of lambda terms;
wenzelm
parents:
diff changeset
    16
sig
43779
47bec02c6762 more uniform Term and Term_XML modules;
wenzelm
parents: 43778
diff changeset
    17
  structure Encode: TERM_XML_OPS
47bec02c6762 more uniform Term and Term_XML modules;
wenzelm
parents: 43778
diff changeset
    18
  structure Decode: TERM_XML_OPS
43729
07d3c6afa865 XML data representation of lambda terms;
wenzelm
parents:
diff changeset
    19
end;
07d3c6afa865 XML data representation of lambda terms;
wenzelm
parents:
diff changeset
    20
07d3c6afa865 XML data representation of lambda terms;
wenzelm
parents:
diff changeset
    21
structure Term_XML: TERM_XML =
07d3c6afa865 XML data representation of lambda terms;
wenzelm
parents:
diff changeset
    22
struct
07d3c6afa865 XML data representation of lambda terms;
wenzelm
parents:
diff changeset
    23
43731
70072780e095 inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents: 43730
diff changeset
    24
structure Encode =
43729
07d3c6afa865 XML data representation of lambda terms;
wenzelm
parents:
diff changeset
    25
struct
07d3c6afa865 XML data representation of lambda terms;
wenzelm
parents:
diff changeset
    26
43767
e0219ef7f84c tuned XML modules;
wenzelm
parents: 43731
diff changeset
    27
open XML.Encode;
43729
07d3c6afa865 XML data representation of lambda terms;
wenzelm
parents:
diff changeset
    28
43730
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents: 43729
diff changeset
    29
val sort = list string;
43729
07d3c6afa865 XML data representation of lambda terms;
wenzelm
parents:
diff changeset
    30
07d3c6afa865 XML data representation of lambda terms;
wenzelm
parents:
diff changeset
    31
fun typ T = T |> variant
43789
wenzelm
parents: 43779
diff changeset
    32
 [fn Type (a, b) => ([a], list typ b),
43778
ce9189450447 more compact representation of XML data (notably sort/typ/term), using properties as vector of atomic values;
wenzelm
parents: 43767
diff changeset
    33
  fn TFree (a, b) => ([a], sort b),
ce9189450447 more compact representation of XML data (notably sort/typ/term), using properties as vector of atomic values;
wenzelm
parents: 43767
diff changeset
    34
  fn TVar ((a, b), c) => ([a, int_atom b], sort c)];
43729
07d3c6afa865 XML data representation of lambda terms;
wenzelm
parents:
diff changeset
    35
07d3c6afa865 XML data representation of lambda terms;
wenzelm
parents:
diff changeset
    36
fun term t = t |> variant
43778
ce9189450447 more compact representation of XML data (notably sort/typ/term), using properties as vector of atomic values;
wenzelm
parents: 43767
diff changeset
    37
 [fn Const (a, b) => ([a], typ b),
ce9189450447 more compact representation of XML data (notably sort/typ/term), using properties as vector of atomic values;
wenzelm
parents: 43767
diff changeset
    38
  fn Free (a, b) => ([a], typ b),
ce9189450447 more compact representation of XML data (notably sort/typ/term), using properties as vector of atomic values;
wenzelm
parents: 43767
diff changeset
    39
  fn Var ((a, b), c) => ([a, int_atom b], typ c),
ce9189450447 more compact representation of XML data (notably sort/typ/term), using properties as vector of atomic values;
wenzelm
parents: 43767
diff changeset
    40
  fn Bound a => ([int_atom a], []),
ce9189450447 more compact representation of XML data (notably sort/typ/term), using properties as vector of atomic values;
wenzelm
parents: 43767
diff changeset
    41
  fn Abs (a, b, c) => ([a], pair typ term (b, c)),
43789
wenzelm
parents: 43779
diff changeset
    42
  fn op $ a => ([], pair term term a)];
43729
07d3c6afa865 XML data representation of lambda terms;
wenzelm
parents:
diff changeset
    43
07d3c6afa865 XML data representation of lambda terms;
wenzelm
parents:
diff changeset
    44
end;
07d3c6afa865 XML data representation of lambda terms;
wenzelm
parents:
diff changeset
    45
43731
70072780e095 inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents: 43730
diff changeset
    46
structure Decode =
43729
07d3c6afa865 XML data representation of lambda terms;
wenzelm
parents:
diff changeset
    47
struct
07d3c6afa865 XML data representation of lambda terms;
wenzelm
parents:
diff changeset
    48
43767
e0219ef7f84c tuned XML modules;
wenzelm
parents: 43731
diff changeset
    49
open XML.Decode;
43729
07d3c6afa865 XML data representation of lambda terms;
wenzelm
parents:
diff changeset
    50
43730
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents: 43729
diff changeset
    51
val sort = list string;
43729
07d3c6afa865 XML data representation of lambda terms;
wenzelm
parents:
diff changeset
    52
07d3c6afa865 XML data representation of lambda terms;
wenzelm
parents:
diff changeset
    53
fun typ T = T |> variant
43778
ce9189450447 more compact representation of XML data (notably sort/typ/term), using properties as vector of atomic values;
wenzelm
parents: 43767
diff changeset
    54
 [fn ([a], b) => Type (a, list typ b),
ce9189450447 more compact representation of XML data (notably sort/typ/term), using properties as vector of atomic values;
wenzelm
parents: 43767
diff changeset
    55
  fn ([a], b) => TFree (a, sort b),
ce9189450447 more compact representation of XML data (notably sort/typ/term), using properties as vector of atomic values;
wenzelm
parents: 43767
diff changeset
    56
  fn ([a, b], c) => TVar ((a, int_atom b), sort c)];
43729
07d3c6afa865 XML data representation of lambda terms;
wenzelm
parents:
diff changeset
    57
07d3c6afa865 XML data representation of lambda terms;
wenzelm
parents:
diff changeset
    58
fun term t = t |> variant
43778
ce9189450447 more compact representation of XML data (notably sort/typ/term), using properties as vector of atomic values;
wenzelm
parents: 43767
diff changeset
    59
 [fn ([a], b) => Const (a, typ b),
ce9189450447 more compact representation of XML data (notably sort/typ/term), using properties as vector of atomic values;
wenzelm
parents: 43767
diff changeset
    60
  fn ([a], b) => Free (a, typ b),
ce9189450447 more compact representation of XML data (notably sort/typ/term), using properties as vector of atomic values;
wenzelm
parents: 43767
diff changeset
    61
  fn ([a, b], c) => Var ((a, int_atom b), typ c),
ce9189450447 more compact representation of XML data (notably sort/typ/term), using properties as vector of atomic values;
wenzelm
parents: 43767
diff changeset
    62
  fn ([a], []) => Bound (int_atom a),
ce9189450447 more compact representation of XML data (notably sort/typ/term), using properties as vector of atomic values;
wenzelm
parents: 43767
diff changeset
    63
  fn ([a], b) => let val (c, d) = pair typ term b in Abs (a, c, d) end,
ce9189450447 more compact representation of XML data (notably sort/typ/term), using properties as vector of atomic values;
wenzelm
parents: 43767
diff changeset
    64
  fn ([], a) => op $ (pair term term a)];
43729
07d3c6afa865 XML data representation of lambda terms;
wenzelm
parents:
diff changeset
    65
07d3c6afa865 XML data representation of lambda terms;
wenzelm
parents:
diff changeset
    66
end;
07d3c6afa865 XML data representation of lambda terms;
wenzelm
parents:
diff changeset
    67
07d3c6afa865 XML data representation of lambda terms;
wenzelm
parents:
diff changeset
    68
end;