src/Pure/term.scala
author wenzelm
Wed, 22 Apr 2020 17:52:14 +0200
changeset 71781 3fd54f7f52b0
parent 71777 3875815f5967
child 73024 337e1b135d2f
permissions -rw-r--r--
tuned signature -- avoid warnings;
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
70852
ee2f490a06b4 clarified signature default;
wenzelm
parents: 70851
diff changeset
    16
  sealed case class Indexname(name: String, index: Int = 0)
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
70850
006214c581ee more operations for type classes;
wenzelm
parents: 70846
diff changeset
    33
  type Class = String
006214c581ee more operations for type classes;
wenzelm
parents: 70846
diff changeset
    34
  type Sort = List[Class]
43730
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
70846
3777d9aaaaad clarified output and input of Typ/Term;
wenzelm
parents: 70834
diff changeset
    38
  {
3777d9aaaaad clarified output and input of Typ/Term;
wenzelm
parents: 70834
diff changeset
    39
    override def toString: String =
3777d9aaaaad clarified output and input of Typ/Term;
wenzelm
parents: 70834
diff changeset
    40
      if (this == dummyT) "_"
3777d9aaaaad clarified output and input of Typ/Term;
wenzelm
parents: 70834
diff changeset
    41
      else "Type(" + name + (if (args.isEmpty) "" else "," + args) + ")"
3777d9aaaaad clarified output and input of Typ/Term;
wenzelm
parents: 70834
diff changeset
    42
  }
70851
6d01eca49b34 clarified signature default;
wenzelm
parents: 70850
diff changeset
    43
  case class TFree(name: String, sort: Sort = Nil) extends Typ
70846
3777d9aaaaad clarified output and input of Typ/Term;
wenzelm
parents: 70834
diff changeset
    44
  {
3777d9aaaaad clarified output and input of Typ/Term;
wenzelm
parents: 70834
diff changeset
    45
    override def toString: String =
70851
6d01eca49b34 clarified signature default;
wenzelm
parents: 70850
diff changeset
    46
      "TFree(" + name + (if (sort.isEmpty) "" else "," + sort) + ")"
70846
3777d9aaaaad clarified output and input of Typ/Term;
wenzelm
parents: 70834
diff changeset
    47
  }
70851
6d01eca49b34 clarified signature default;
wenzelm
parents: 70850
diff changeset
    48
  case class TVar(name: Indexname, sort: Sort = Nil) extends Typ
70846
3777d9aaaaad clarified output and input of Typ/Term;
wenzelm
parents: 70834
diff changeset
    49
  {
3777d9aaaaad clarified output and input of Typ/Term;
wenzelm
parents: 70834
diff changeset
    50
    override def toString: String =
70851
6d01eca49b34 clarified signature default;
wenzelm
parents: 70850
diff changeset
    51
      "TVar(" + name + (if (sort.isEmpty) "" else "," + sort) + ")"
70846
3777d9aaaaad clarified output and input of Typ/Term;
wenzelm
parents: 70834
diff changeset
    52
  }
71601
97ccf48c2f0c misc tuning based on hints by IntelliJ IDEA;
wenzelm
parents: 71221
diff changeset
    53
  val dummyT: Type = Type("dummy")
43730
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    54
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    55
  sealed abstract class Term
71221
4dfb7c937126 more operations;
wenzelm
parents: 70916
diff changeset
    56
  {
4dfb7c937126 more operations;
wenzelm
parents: 70916
diff changeset
    57
    def head: Term =
4dfb7c937126 more operations;
wenzelm
parents: 70916
diff changeset
    58
      this match {
4dfb7c937126 more operations;
wenzelm
parents: 70916
diff changeset
    59
        case App(fun, _) => fun.head
4dfb7c937126 more operations;
wenzelm
parents: 70916
diff changeset
    60
        case _ => this
4dfb7c937126 more operations;
wenzelm
parents: 70916
diff changeset
    61
      }
4dfb7c937126 more operations;
wenzelm
parents: 70916
diff changeset
    62
  }
70846
3777d9aaaaad clarified output and input of Typ/Term;
wenzelm
parents: 70834
diff changeset
    63
  case class Const(name: String, typargs: List[Typ] = Nil) extends Term
3777d9aaaaad clarified output and input of Typ/Term;
wenzelm
parents: 70834
diff changeset
    64
  {
3777d9aaaaad clarified output and input of Typ/Term;
wenzelm
parents: 70834
diff changeset
    65
    override def toString: String =
70897
2cc7c05b3b3c support dummy term;
wenzelm
parents: 70852
diff changeset
    66
      if (this == dummy) "_"
2cc7c05b3b3c support dummy term;
wenzelm
parents: 70852
diff changeset
    67
      else "Const(" + name + (if (typargs.isEmpty) "" else "," + typargs) + ")"
70846
3777d9aaaaad clarified output and input of Typ/Term;
wenzelm
parents: 70834
diff changeset
    68
  }
43730
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    69
  case class Free(name: String, typ: Typ = dummyT) extends Term
70846
3777d9aaaaad clarified output and input of Typ/Term;
wenzelm
parents: 70834
diff changeset
    70
  {
3777d9aaaaad clarified output and input of Typ/Term;
wenzelm
parents: 70834
diff changeset
    71
    override def toString: String =
3777d9aaaaad clarified output and input of Typ/Term;
wenzelm
parents: 70834
diff changeset
    72
      "Free(" + name + (if (typ == dummyT) "" else "," + typ) + ")"
3777d9aaaaad clarified output and input of Typ/Term;
wenzelm
parents: 70834
diff changeset
    73
  }
43730
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    74
  case class Var(name: Indexname, typ: Typ = dummyT) extends Term
70846
3777d9aaaaad clarified output and input of Typ/Term;
wenzelm
parents: 70834
diff changeset
    75
  {
3777d9aaaaad clarified output and input of Typ/Term;
wenzelm
parents: 70834
diff changeset
    76
    override def toString: String =
3777d9aaaaad clarified output and input of Typ/Term;
wenzelm
parents: 70834
diff changeset
    77
      "Var(" + name + (if (typ == dummyT) "" else "," + typ) + ")"
3777d9aaaaad clarified output and input of Typ/Term;
wenzelm
parents: 70834
diff changeset
    78
  }
43730
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    79
  case class Bound(index: Int) extends Term
70846
3777d9aaaaad clarified output and input of Typ/Term;
wenzelm
parents: 70834
diff changeset
    80
  case class Abs(name: String, typ: Typ, body: Term) extends Term
43730
a0ed7bc688b5 lambda terms with XML data representation in Scala;
wenzelm
parents:
diff changeset
    81
  case class App(fun: Term, arg: Term) extends Term
70850
006214c581ee more operations for type classes;
wenzelm
parents: 70846
diff changeset
    82
  {
006214c581ee more operations for type classes;
wenzelm
parents: 70846
diff changeset
    83
    override def toString: String =
006214c581ee more operations for type classes;
wenzelm
parents: 70846
diff changeset
    84
      this match {
006214c581ee more operations for type classes;
wenzelm
parents: 70846
diff changeset
    85
        case OFCLASS(ty, c) => "OFCLASS(" + ty + "," + c + ")"
006214c581ee more operations for type classes;
wenzelm
parents: 70846
diff changeset
    86
        case _ => "App(" + fun + "," + arg + ")"
006214c581ee more operations for type classes;
wenzelm
parents: 70846
diff changeset
    87
      }
006214c581ee more operations for type classes;
wenzelm
parents: 70846
diff changeset
    88
  }
68265
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
    89
70897
2cc7c05b3b3c support dummy term;
wenzelm
parents: 70852
diff changeset
    90
  def dummy_pattern(ty: Typ): Term = Const("Pure.dummy_pattern", List(ty))
2cc7c05b3b3c support dummy term;
wenzelm
parents: 70852
diff changeset
    91
  val dummy: Term = dummy_pattern(dummyT)
2cc7c05b3b3c support dummy term;
wenzelm
parents: 70852
diff changeset
    92
70533
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
    93
  sealed abstract class Proof
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
    94
  case object MinProof extends Proof
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
    95
  case class PBound(index: Int) extends Proof
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
    96
  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
    97
  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
    98
  case class Appt(fun: Proof, arg: Term) extends Proof
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
    99
  case class AppP(fun: Proof, arg: Proof) extends Proof
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   100
  case class Hyp(hyp: Term) extends Proof
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   101
  case class PAxm(name: String, types: List[Typ]) extends Proof
71777
3875815f5967 clarified signature: avoid clash with Isabelle/Scala Term.OFCLASS on case-insensible file-system;
wenzelm
parents: 71601
diff changeset
   102
  case class PClass(typ: Typ, cls: Class) extends Proof
70834
614ca81fa28e clarified oracle_proof;
wenzelm
parents: 70786
diff changeset
   103
  case class Oracle(name: String, prop: Term, types: List[Typ]) extends Proof
70916
4c15217d6266 clarified signature: name of standard_proof is authentic, otherwise empty;
wenzelm
parents: 70897
diff changeset
   104
  case class PThm(serial: Long, theory_name: String, name: String, types: List[Typ]) extends Proof
70533
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   105
68265
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   106
70850
006214c581ee more operations for type classes;
wenzelm
parents: 70846
diff changeset
   107
  /* type classes within the logic */
70385
68d2c533db9c more operations: support type classes within the logic;
wenzelm
parents: 70383
diff changeset
   108
70850
006214c581ee more operations for type classes;
wenzelm
parents: 70846
diff changeset
   109
  object Class_Const
006214c581ee more operations for type classes;
wenzelm
parents: 70846
diff changeset
   110
  {
006214c581ee more operations for type classes;
wenzelm
parents: 70846
diff changeset
   111
    val suffix = "_class"
006214c581ee more operations for type classes;
wenzelm
parents: 70846
diff changeset
   112
    def apply(c: Class): String = c + suffix
006214c581ee more operations for type classes;
wenzelm
parents: 70846
diff changeset
   113
    def unapply(s: String): Option[Class] =
006214c581ee more operations for type classes;
wenzelm
parents: 70846
diff changeset
   114
      if (s.endsWith(suffix)) Some(s.substring(0, s.length - suffix.length)) else None
006214c581ee more operations for type classes;
wenzelm
parents: 70846
diff changeset
   115
  }
70385
68d2c533db9c more operations: support type classes within the logic;
wenzelm
parents: 70383
diff changeset
   116
70850
006214c581ee more operations for type classes;
wenzelm
parents: 70846
diff changeset
   117
  object OFCLASS
006214c581ee more operations for type classes;
wenzelm
parents: 70846
diff changeset
   118
  {
006214c581ee more operations for type classes;
wenzelm
parents: 70846
diff changeset
   119
    def apply(ty: Typ, s: Sort): List[Term] = s.map(c => apply(ty, c))
006214c581ee more operations for type classes;
wenzelm
parents: 70846
diff changeset
   120
006214c581ee more operations for type classes;
wenzelm
parents: 70846
diff changeset
   121
    def apply(ty: Typ, c: Class): Term =
006214c581ee more operations for type classes;
wenzelm
parents: 70846
diff changeset
   122
      App(Const(Class_Const(c), List(ty)), Const(Pure_Thy.TYPE, List(ty)))
70385
68d2c533db9c more operations: support type classes within the logic;
wenzelm
parents: 70383
diff changeset
   123
70850
006214c581ee more operations for type classes;
wenzelm
parents: 70846
diff changeset
   124
    def unapply(t: Term): Option[(Typ, String)] =
006214c581ee more operations for type classes;
wenzelm
parents: 70846
diff changeset
   125
      t match {
006214c581ee more operations for type classes;
wenzelm
parents: 70846
diff changeset
   126
        case App(Const(Class_Const(c), List(ty)), Const(Pure_Thy.TYPE, List(ty1)))
006214c581ee more operations for type classes;
wenzelm
parents: 70846
diff changeset
   127
        if ty == ty1 => Some((ty, c))
006214c581ee more operations for type classes;
wenzelm
parents: 70846
diff changeset
   128
        case _ => None
006214c581ee more operations for type classes;
wenzelm
parents: 70846
diff changeset
   129
      }
70385
68d2c533db9c more operations: support type classes within the logic;
wenzelm
parents: 70383
diff changeset
   130
  }
68d2c533db9c more operations: support type classes within the logic;
wenzelm
parents: 70383
diff changeset
   131
68d2c533db9c more operations: support type classes within the logic;
wenzelm
parents: 70383
diff changeset
   132
68265
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   133
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   134
  /** cache **/
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   135
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   136
  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
   137
    new Cache(initial_size, max_string)
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
  class Cache private[Term](initial_size: Int, max_string: Int)
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   140
    extends isabelle.Cache(initial_size, max_string)
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   141
  {
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   142
    protected def cache_indexname(x: Indexname): Indexname =
70536
fe4d545f12e3 clarified type Indexname, with plain value Int;
wenzelm
parents: 70535
diff changeset
   143
      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
   144
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   145
    protected def cache_sort(x: Sort): Sort =
71601
97ccf48c2f0c misc tuning based on hints by IntelliJ IDEA;
wenzelm
parents: 71221
diff changeset
   146
      if (x.isEmpty) Nil else lookup(x) getOrElse store(x.map(cache_string))
68265
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   147
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   148
    protected def cache_typ(x: Typ): Typ =
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   149
    {
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   150
      if (x == dummyT) dummyT
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   151
      else
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   152
        lookup(x) match {
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   153
          case Some(y) => y
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   154
          case None =>
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   155
            x match {
70533
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   156
              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
   157
              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
   158
              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
   159
            }
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   160
        }
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   161
    }
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   162
70533
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   163
    protected def cache_typs(x: List[Typ]): List[Typ] =
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   164
    {
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   165
      if (x.isEmpty) Nil
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   166
      else {
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   167
        lookup(x) match {
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   168
          case Some(y) => y
71601
97ccf48c2f0c misc tuning based on hints by IntelliJ IDEA;
wenzelm
parents: 71221
diff changeset
   169
          case None => store(x.map(cache_typ))
70533
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   170
        }
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   171
      }
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   172
    }
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   173
68265
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   174
    protected def cache_term(x: Term): Term =
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   175
    {
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   176
      lookup(x) match {
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   177
        case Some(y) => y
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   178
        case None =>
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   179
          x match {
70784
799437173553 Term_XML.Encode/Decode.term uses Const "typargs";
wenzelm
parents: 70559
diff changeset
   180
            case Const(name, typargs) => store(Const(cache_string(name), cache_typs(typargs)))
68265
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   181
            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
   182
            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
   183
            case Bound(_) => store(x)
68265
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   184
            case Abs(name, typ, body) =>
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   185
              store(Abs(cache_string(name), cache_typ(typ), cache_term(body)))
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   186
            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
   187
          }
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   188
      }
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   189
    }
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   190
70533
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   191
    protected def cache_proof(x: Proof): Proof =
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
      if (x == MinProof) MinProof
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   194
      else {
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   195
        lookup(x) match {
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   196
          case Some(y) => y
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   197
          case None =>
71781
3fd54f7f52b0 tuned signature -- avoid warnings;
wenzelm
parents: 71777
diff changeset
   198
            (x: @unchecked) match {
70536
fe4d545f12e3 clarified type Indexname, with plain value Int;
wenzelm
parents: 70535
diff changeset
   199
              case PBound(_) => store(x)
70533
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   200
              case Abst(name, typ, body) =>
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   201
                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
   202
              case AbsP(name, hyp, body) =>
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   203
                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
   204
              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
   205
              case AppP(fun, arg) => store(AppP(cache_proof(fun), cache_proof(arg)))
70535
de6062ac70b6 more complete pattern match;
wenzelm
parents: 70533
diff changeset
   206
              case Hyp(hyp) => store(Hyp(cache_term(hyp)))
de6062ac70b6 more complete pattern match;
wenzelm
parents: 70533
diff changeset
   207
              case PAxm(name, types) => store(PAxm(cache_string(name), cache_typs(types)))
71777
3875815f5967 clarified signature: avoid clash with Isabelle/Scala Term.OFCLASS on case-insensible file-system;
wenzelm
parents: 71601
diff changeset
   208
              case PClass(typ, cls) => store(PClass(cache_typ(typ), cache_string(cls)))
70533
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   209
              case Oracle(name, prop, types) =>
70834
614ca81fa28e clarified oracle_proof;
wenzelm
parents: 70786
diff changeset
   210
                store(Oracle(cache_string(name), cache_term(prop), cache_typs(types)))
70538
fc9ba6fe367f clarified PThm: theory_name simplifies retrieval from exports;
wenzelm
parents: 70537
diff changeset
   211
              case PThm(serial, theory_name, name, types) =>
fc9ba6fe367f clarified PThm: theory_name simplifies retrieval from exports;
wenzelm
parents: 70537
diff changeset
   212
                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
   213
            }
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   214
        }
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   215
      }
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   216
    }
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   217
68265
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   218
    // main methods
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   219
    def indexname(x: Indexname): Indexname = synchronized { cache_indexname(x) }
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   220
    def sort(x: Sort): Sort = synchronized { cache_sort(x) }
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   221
    def typ(x: Typ): Typ = synchronized { cache_typ(x) }
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   222
    def term(x: Term): Term = synchronized { cache_term(x) }
70533
031620901fcd support for (fully reconstructed) proof terms in Scala;
wenzelm
parents: 70385
diff changeset
   223
    def proof(x: Proof): Proof = synchronized { cache_proof(x) }
68265
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   224
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   225
    def position(x: Position.T): Position.T =
f0899dad4877 more general cache, also for term substructures;
wenzelm
parents: 43779
diff changeset
   226
      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
   227
  }
43731
70072780e095 inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents: 43730
diff changeset
   228
}