src/Pure/term.scala
author wenzelm
Sat, 17 Aug 2019 11:52:47 +0200
changeset 70559 c92443e8d724
parent 70554 10d41bf28b92
child 70784 799437173553
permissions -rw-r--r--
clarified type for recorded oracles;
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
43779
47bec02c6762 more uniform Term and Term_XML modules;
wenzelm
parents: 43778
diff changeset
     4
Lambda terms, types, sorts.
43730
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
{
70383
38ac2e714729 more operations (avoid clones in Isabelle/MMT and Isabelle/Dedukti);
wenzelm
parents: 68265
diff changeset
    14
  /* types and terms */
38ac2e714729 more operations (avoid clones in Isabelle/MMT and Isabelle/Dedukti);
wenzelm
parents: 68265
diff changeset
    15
70536
fe4d545f12e3 clarified type Indexname, with plain value Int;
wenzelm
parents: 70535
diff changeset
    16
  sealed case class Indexname(name: String, index: Int)
70537
17160e0a60b6 Indexname.toString according to string_of_vname' in ML;
wenzelm
parents: 70536
diff changeset
    17
  {
17160e0a60b6 Indexname.toString according to string_of_vname' in ML;
wenzelm
parents: 70536
diff changeset
    18
    override def toString: String =
17160e0a60b6 Indexname.toString according to string_of_vname' in ML;
wenzelm
parents: 70536
diff changeset
    19
      if (index == -1) name
17160e0a60b6 Indexname.toString according to string_of_vname' in ML;
wenzelm
parents: 70536
diff changeset
    20
      else {
17160e0a60b6 Indexname.toString according to string_of_vname' in ML;
wenzelm
parents: 70536
diff changeset
    21
        val dot =
17160e0a60b6 Indexname.toString according to string_of_vname' in ML;
wenzelm
parents: 70536
diff changeset
    22
          Symbol.explode(name).reverse match {
17160e0a60b6 Indexname.toString according to string_of_vname' in ML;
wenzelm
parents: 70536
diff changeset
    23
            case _ :: s :: _ if s == Symbol.sub_decoded || s == Symbol.sub => false
17160e0a60b6 Indexname.toString according to string_of_vname' in ML;
wenzelm
parents: 70536
diff changeset
    24
            case c :: _ => Symbol.is_digit(c)
17160e0a60b6 Indexname.toString according to string_of_vname' in ML;
wenzelm
parents: 70536
diff changeset
    25
            case _ => true
17160e0a60b6 Indexname.toString according to string_of_vname' in ML;
wenzelm
parents: 70536
diff changeset
    26
          }
17160e0a60b6 Indexname.toString according to string_of_vname' in ML;
wenzelm
parents: 70536
diff changeset
    27
        if (dot) "?" + name + "." + index
17160e0a60b6 Indexname.toString according to string_of_vname' in ML;
wenzelm
parents: 70536
diff changeset
    28
        else if (index != 0) "?" + name + index
17160e0a60b6 Indexname.toString according to string_of_vname' in ML;
wenzelm
parents: 70536
diff changeset
    29
        else "?" + name
17160e0a60b6 Indexname.toString according to string_of_vname' in ML;
wenzelm
parents: 70536
diff changeset
    30
      }
17160e0a60b6 Indexname.toString according to string_of_vname' in ML;
wenzelm
parents: 70536
diff changeset
    31
  }
43730
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    32
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    33
  type Sort = List[String]
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    34
  val dummyS: Sort = List("")
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
  sealed abstract class Typ
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    37
  case class Type(name: String, args: List[Typ] = Nil) extends Typ
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    38
  case class TFree(name: String, sort: Sort = dummyS) extends Typ
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    39
  case class TVar(name: Indexname, sort: Sort = dummyS) extends Typ
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    40
  val dummyT = Type("dummy")
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
  sealed abstract class Term
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    43
  case class Const(name: String, typ: Typ = dummyT) extends Term
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    44
  case class Free(name: String, typ: Typ = dummyT) extends Term
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    45
  case class Var(name: Indexname, typ: Typ = dummyT) extends Term
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    46
  case class Bound(index: Int) extends Term
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    47
  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
    48
  case class App(fun: Term, arg: Term) extends Term
68265
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
    49
70533
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
    50
  sealed abstract class Proof
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
    51
  case object MinProof extends Proof
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
    52
  case class PBound(index: Int) extends Proof
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
    53
  case class Abst(name: String, typ: Typ, body: Proof) extends Proof
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
    54
  case class AbsP(name: String, hyp: Term, body: Proof) extends Proof
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
    55
  case class Appt(fun: Proof, arg: Term) extends Proof
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
    56
  case class AppP(fun: Proof, arg: Proof) extends Proof
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
    57
  case class Hyp(hyp: Term) extends Proof
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
    58
  case class PAxm(name: String, types: List[Typ]) extends Proof
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
    59
  case class OfClass(typ: Typ, cls: String) extends Proof
70559
c92443e8d724 clarified type for recorded oracles;
wenzelm
parents: 70554
diff changeset
    60
  case class Oracle(name: String, prop: Option[Term], types: List[Typ]) extends Proof
70554
10d41bf28b92 clarified signature;
wenzelm
parents: 70538
diff changeset
    61
  case class PThm(serial: Long, theory_name: String, approximative_name: String, types: List[Typ])
70538
fc9ba6fe367f clarified PThm: theory_name simplifies retrieval from exports;
wenzelm
parents: 70537
diff changeset
    62
    extends Proof
70533
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
    63
68265
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
    64
70385
68d2c533db9c more operations: support type classes within the logic;
wenzelm
parents: 70383
diff changeset
    65
  /* Pure logic */
68d2c533db9c more operations: support type classes within the logic;
wenzelm
parents: 70383
diff changeset
    66
68d2c533db9c more operations: support type classes within the logic;
wenzelm
parents: 70383
diff changeset
    67
  def itselfT(ty: Typ): Typ = Type(Pure_Thy.ITSELF, List(ty))
68d2c533db9c more operations: support type classes within the logic;
wenzelm
parents: 70383
diff changeset
    68
  val propT: Typ = Type(Pure_Thy.PROP, Nil)
68d2c533db9c more operations: support type classes within the logic;
wenzelm
parents: 70383
diff changeset
    69
  def funT(ty1: Typ, ty2: Typ): Typ = Type(Pure_Thy.FUN, List(ty1, ty2))
68d2c533db9c more operations: support type classes within the logic;
wenzelm
parents: 70383
diff changeset
    70
68d2c533db9c more operations: support type classes within the logic;
wenzelm
parents: 70383
diff changeset
    71
  def mk_type(ty: Typ): Term = Const(Pure_Thy.TYPE, itselfT(ty))
68d2c533db9c more operations: support type classes within the logic;
wenzelm
parents: 70383
diff changeset
    72
68d2c533db9c more operations: support type classes within the logic;
wenzelm
parents: 70383
diff changeset
    73
  def const_of_class(c: String): String = c + "_class"
68d2c533db9c more operations: support type classes within the logic;
wenzelm
parents: 70383
diff changeset
    74
68d2c533db9c more operations: support type classes within the logic;
wenzelm
parents: 70383
diff changeset
    75
  def mk_of_sort(ty: Typ, s: Sort): List[Term] =
68d2c533db9c more operations: support type classes within the logic;
wenzelm
parents: 70383
diff changeset
    76
  {
68d2c533db9c more operations: support type classes within the logic;
wenzelm
parents: 70383
diff changeset
    77
    val class_type = funT(itselfT(ty), propT)
68d2c533db9c more operations: support type classes within the logic;
wenzelm
parents: 70383
diff changeset
    78
    val t = mk_type(ty)
68d2c533db9c more operations: support type classes within the logic;
wenzelm
parents: 70383
diff changeset
    79
    s.map(c => App(Const(const_of_class(c), class_type), t))
68d2c533db9c more operations: support type classes within the logic;
wenzelm
parents: 70383
diff changeset
    80
  }
68d2c533db9c more operations: support type classes within the logic;
wenzelm
parents: 70383
diff changeset
    81
68d2c533db9c more operations: support type classes within the logic;
wenzelm
parents: 70383
diff changeset
    82
70383
38ac2e714729 more operations (avoid clones in Isabelle/MMT and Isabelle/Dedukti);
wenzelm
parents: 68265
diff changeset
    83
  /* type arguments of consts */
38ac2e714729 more operations (avoid clones in Isabelle/MMT and Isabelle/Dedukti);
wenzelm
parents: 68265
diff changeset
    84
38ac2e714729 more operations (avoid clones in Isabelle/MMT and Isabelle/Dedukti);
wenzelm
parents: 68265
diff changeset
    85
  def const_typargs(name: String, typ: Typ, typargs: List[String], decl: Typ): List[Typ] =
38ac2e714729 more operations (avoid clones in Isabelle/MMT and Isabelle/Dedukti);
wenzelm
parents: 68265
diff changeset
    86
  {
38ac2e714729 more operations (avoid clones in Isabelle/MMT and Isabelle/Dedukti);
wenzelm
parents: 68265
diff changeset
    87
    var subst = Map.empty[String, Typ]
38ac2e714729 more operations (avoid clones in Isabelle/MMT and Isabelle/Dedukti);
wenzelm
parents: 68265
diff changeset
    88
38ac2e714729 more operations (avoid clones in Isabelle/MMT and Isabelle/Dedukti);
wenzelm
parents: 68265
diff changeset
    89
    def bad_match(): Nothing = error("Malformed type instance for " + name + ": " + typ)
38ac2e714729 more operations (avoid clones in Isabelle/MMT and Isabelle/Dedukti);
wenzelm
parents: 68265
diff changeset
    90
    def raw_match(arg: (Typ, Typ))
38ac2e714729 more operations (avoid clones in Isabelle/MMT and Isabelle/Dedukti);
wenzelm
parents: 68265
diff changeset
    91
    {
38ac2e714729 more operations (avoid clones in Isabelle/MMT and Isabelle/Dedukti);
wenzelm
parents: 68265
diff changeset
    92
      arg match {
38ac2e714729 more operations (avoid clones in Isabelle/MMT and Isabelle/Dedukti);
wenzelm
parents: 68265
diff changeset
    93
        case (TFree(a, _), ty) =>
38ac2e714729 more operations (avoid clones in Isabelle/MMT and Isabelle/Dedukti);
wenzelm
parents: 68265
diff changeset
    94
          subst.get(a) match {
38ac2e714729 more operations (avoid clones in Isabelle/MMT and Isabelle/Dedukti);
wenzelm
parents: 68265
diff changeset
    95
            case None => subst += (a -> ty)
38ac2e714729 more operations (avoid clones in Isabelle/MMT and Isabelle/Dedukti);
wenzelm
parents: 68265
diff changeset
    96
            case Some(ty1) => if (ty != ty1) bad_match()
38ac2e714729 more operations (avoid clones in Isabelle/MMT and Isabelle/Dedukti);
wenzelm
parents: 68265
diff changeset
    97
          }
38ac2e714729 more operations (avoid clones in Isabelle/MMT and Isabelle/Dedukti);
wenzelm
parents: 68265
diff changeset
    98
        case (Type(c1, args1), Type(c2, args2)) if c1 == c2 =>
38ac2e714729 more operations (avoid clones in Isabelle/MMT and Isabelle/Dedukti);
wenzelm
parents: 68265
diff changeset
    99
          (args1 zip args2).foreach(raw_match)
38ac2e714729 more operations (avoid clones in Isabelle/MMT and Isabelle/Dedukti);
wenzelm
parents: 68265
diff changeset
   100
        case _ => bad_match()
38ac2e714729 more operations (avoid clones in Isabelle/MMT and Isabelle/Dedukti);
wenzelm
parents: 68265
diff changeset
   101
      }
38ac2e714729 more operations (avoid clones in Isabelle/MMT and Isabelle/Dedukti);
wenzelm
parents: 68265
diff changeset
   102
    }
38ac2e714729 more operations (avoid clones in Isabelle/MMT and Isabelle/Dedukti);
wenzelm
parents: 68265
diff changeset
   103
    raw_match(decl, typ)
38ac2e714729 more operations (avoid clones in Isabelle/MMT and Isabelle/Dedukti);
wenzelm
parents: 68265
diff changeset
   104
38ac2e714729 more operations (avoid clones in Isabelle/MMT and Isabelle/Dedukti);
wenzelm
parents: 68265
diff changeset
   105
    typargs.map(subst)
38ac2e714729 more operations (avoid clones in Isabelle/MMT and Isabelle/Dedukti);
wenzelm
parents: 68265
diff changeset
   106
  }
38ac2e714729 more operations (avoid clones in Isabelle/MMT and Isabelle/Dedukti);
wenzelm
parents: 68265
diff changeset
   107
38ac2e714729 more operations (avoid clones in Isabelle/MMT and Isabelle/Dedukti);
wenzelm
parents: 68265
diff changeset
   108
68265
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   109
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   110
  /** cache **/
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   111
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   112
  def make_cache(initial_size: Int = 131071, max_string: Int = Integer.MAX_VALUE): Cache =
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   113
    new Cache(initial_size, max_string)
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   114
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   115
  class Cache private[Term](initial_size: Int, max_string: Int)
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   116
    extends isabelle.Cache(initial_size, max_string)
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   117
  {
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   118
    protected def cache_indexname(x: Indexname): Indexname =
70536
fe4d545f12e3 clarified type Indexname, with plain value Int;
wenzelm
parents: 70535
diff changeset
   119
      lookup(x) getOrElse store(Indexname(cache_string(x.name), x.index))
68265
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   120
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   121
    protected def cache_sort(x: Sort): Sort =
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   122
      if (x == dummyS) dummyS
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   123
      else lookup(x) getOrElse store(x.map(cache_string(_)))
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   124
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   125
    protected def cache_typ(x: Typ): Typ =
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   126
    {
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   127
      if (x == dummyT) dummyT
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   128
      else
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   129
        lookup(x) match {
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   130
          case Some(y) => y
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   131
          case None =>
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   132
            x match {
70533
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   133
              case Type(name, args) => store(Type(cache_string(name), cache_typs(args)))
68265
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   134
              case TFree(name, sort) => store(TFree(cache_string(name), cache_sort(sort)))
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   135
              case TVar(name, sort) => store(TVar(cache_indexname(name), cache_sort(sort)))
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   136
            }
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   137
        }
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   138
    }
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   139
70533
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   140
    protected def cache_typs(x: List[Typ]): List[Typ] =
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   141
    {
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   142
      if (x.isEmpty) Nil
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   143
      else {
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   144
        lookup(x) match {
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   145
          case Some(y) => y
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   146
          case None => store(x.map(cache_typ(_)))
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   147
        }
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   148
      }
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   149
    }
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   150
68265
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   151
    protected def cache_term(x: Term): Term =
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   152
    {
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   153
      lookup(x) match {
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   154
        case Some(y) => y
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   155
        case None =>
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   156
          x match {
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   157
            case Const(name, typ) => store(Const(cache_string(name), cache_typ(typ)))
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   158
            case Free(name, typ) => store(Free(cache_string(name), cache_typ(typ)))
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   159
            case Var(name, typ) => store(Var(cache_indexname(name), cache_typ(typ)))
70536
fe4d545f12e3 clarified type Indexname, with plain value Int;
wenzelm
parents: 70535
diff changeset
   160
            case Bound(_) => store(x)
68265
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   161
            case Abs(name, typ, body) =>
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   162
              store(Abs(cache_string(name), cache_typ(typ), cache_term(body)))
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   163
            case App(fun, arg) => store(App(cache_term(fun), cache_term(arg)))
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   164
          }
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   165
      }
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   166
    }
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   167
70533
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   168
    protected def cache_proof(x: Proof): Proof =
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   169
    {
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   170
      if (x == MinProof) MinProof
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   171
      else {
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   172
        lookup(x) match {
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   173
          case Some(y) => y
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   174
          case None =>
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   175
            x match {
70536
fe4d545f12e3 clarified type Indexname, with plain value Int;
wenzelm
parents: 70535
diff changeset
   176
              case PBound(_) => store(x)
70533
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   177
              case Abst(name, typ, body) =>
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   178
                store(Abst(cache_string(name), cache_typ(typ), cache_proof(body)))
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   179
              case AbsP(name, hyp, body) =>
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   180
                store(AbsP(cache_string(name), cache_term(hyp), cache_proof(body)))
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   181
              case Appt(fun, arg) => store(Appt(cache_proof(fun), cache_term(arg)))
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   182
              case AppP(fun, arg) => store(AppP(cache_proof(fun), cache_proof(arg)))
70535
de6062ac70b6 more complete pattern match;
wenzelm
parents: 70533
diff changeset
   183
              case Hyp(hyp) => store(Hyp(cache_term(hyp)))
de6062ac70b6 more complete pattern match;
wenzelm
parents: 70533
diff changeset
   184
              case PAxm(name, types) => store(PAxm(cache_string(name), cache_typs(types)))
70533
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   185
              case OfClass(typ, cls) => store(OfClass(cache_typ(typ), cache_string(cls)))
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   186
              case Oracle(name, prop, types) =>
70559
c92443e8d724 clarified type for recorded oracles;
wenzelm
parents: 70554
diff changeset
   187
                store(Oracle(cache_string(name), prop.map(cache_term(_)), cache_typs(types)))
70538
fc9ba6fe367f clarified PThm: theory_name simplifies retrieval from exports;
wenzelm
parents: 70537
diff changeset
   188
              case PThm(serial, theory_name, name, types) =>
fc9ba6fe367f clarified PThm: theory_name simplifies retrieval from exports;
wenzelm
parents: 70537
diff changeset
   189
                store(PThm(serial, cache_string(theory_name), cache_string(name), cache_typs(types)))
70533
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   190
            }
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   191
        }
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   192
      }
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   193
    }
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   194
68265
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   195
    // main methods
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   196
    def indexname(x: Indexname): Indexname = synchronized { cache_indexname(x) }
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   197
    def sort(x: Sort): Sort = synchronized { cache_sort(x) }
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   198
    def typ(x: Typ): Typ = synchronized { cache_typ(x) }
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   199
    def term(x: Term): Term = synchronized { cache_term(x) }
70533
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   200
    def proof(x: Proof): Proof = synchronized { cache_proof(x) }
68265
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   201
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   202
    def position(x: Position.T): Position.T =
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   203
      synchronized { x.map({ case (a, b) => (cache_string(a), cache_string(b)) }) }
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   204
  }
43731
70072780e095 inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents: 43730
diff changeset
   205
}