| author | desharna | 
| Thu, 07 Mar 2024 23:58:01 +0100 | |
| changeset 79807 | afb26a1dea71 | 
| parent 77368 | 7c57d9586f4c | 
| child 80295 | 8a9588ffc133 | 
| permissions | -rw-r--r-- | 
| 43730 | 1 | /* Title: Pure/term.scala | 
| 2 | Author: Makarius | |
| 3 | ||
| 43779 | 4 | Lambda terms, types, sorts. | 
| 43730 | 5 | |
| 6 | Note: Isabelle/ML is the primary environment for logical operations. | |
| 7 | */ | |
| 8 | ||
| 9 | package isabelle | |
| 10 | ||
| 11 | ||
| 75393 | 12 | object Term {
 | 
| 70383 
38ac2e714729
more operations (avoid clones in Isabelle/MMT and Isabelle/Dedukti);
 wenzelm parents: 
68265diff
changeset | 13 | /* types and terms */ | 
| 
38ac2e714729
more operations (avoid clones in Isabelle/MMT and Isabelle/Dedukti);
 wenzelm parents: 
68265diff
changeset | 14 | |
| 75393 | 15 |   sealed case class Indexname(name: String, index: Int = 0) {
 | 
| 74121 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 16 | private lazy val hash: Int = (name, index).hashCode() | 
| 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 17 | override def hashCode(): Int = hash | 
| 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 18 | |
| 70537 
17160e0a60b6
Indexname.toString according to string_of_vname' in ML;
 wenzelm parents: 
70536diff
changeset | 19 | override def toString: String = | 
| 
17160e0a60b6
Indexname.toString according to string_of_vname' in ML;
 wenzelm parents: 
70536diff
changeset | 20 | if (index == -1) name | 
| 
17160e0a60b6
Indexname.toString according to string_of_vname' in ML;
 wenzelm parents: 
70536diff
changeset | 21 |       else {
 | 
| 
17160e0a60b6
Indexname.toString according to string_of_vname' in ML;
 wenzelm parents: 
70536diff
changeset | 22 | val dot = | 
| 
17160e0a60b6
Indexname.toString according to string_of_vname' in ML;
 wenzelm parents: 
70536diff
changeset | 23 |           Symbol.explode(name).reverse match {
 | 
| 
17160e0a60b6
Indexname.toString according to string_of_vname' in ML;
 wenzelm parents: 
70536diff
changeset | 24 | case _ :: s :: _ if s == Symbol.sub_decoded || s == Symbol.sub => false | 
| 
17160e0a60b6
Indexname.toString according to string_of_vname' in ML;
 wenzelm parents: 
70536diff
changeset | 25 | case c :: _ => Symbol.is_digit(c) | 
| 
17160e0a60b6
Indexname.toString according to string_of_vname' in ML;
 wenzelm parents: 
70536diff
changeset | 26 | case _ => true | 
| 
17160e0a60b6
Indexname.toString according to string_of_vname' in ML;
 wenzelm parents: 
70536diff
changeset | 27 | } | 
| 
17160e0a60b6
Indexname.toString according to string_of_vname' in ML;
 wenzelm parents: 
70536diff
changeset | 28 | if (dot) "?" + name + "." + index | 
| 
17160e0a60b6
Indexname.toString according to string_of_vname' in ML;
 wenzelm parents: 
70536diff
changeset | 29 | else if (index != 0) "?" + name + index | 
| 
17160e0a60b6
Indexname.toString according to string_of_vname' in ML;
 wenzelm parents: 
70536diff
changeset | 30 | else "?" + name | 
| 
17160e0a60b6
Indexname.toString according to string_of_vname' in ML;
 wenzelm parents: 
70536diff
changeset | 31 | } | 
| 
17160e0a60b6
Indexname.toString according to string_of_vname' in ML;
 wenzelm parents: 
70536diff
changeset | 32 | } | 
| 43730 | 33 | |
| 70850 | 34 | type Class = String | 
| 35 | type Sort = List[Class] | |
| 43730 | 36 | |
| 37 | sealed abstract class Typ | |
| 75393 | 38 |   case class Type(name: String, args: List[Typ] = Nil) extends Typ {
 | 
| 74121 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 39 |     private lazy val hash: Int = ("Type", name, args).hashCode()
 | 
| 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 40 | override def hashCode(): Int = hash | 
| 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 41 | |
| 70846 | 42 | override def toString: String = | 
| 43 | if (this == dummyT) "_" | |
| 77368 | 44 |       else "Type(" + name + if_proper(args, "," + args) + ")"
 | 
| 70846 | 45 | } | 
| 75393 | 46 |   case class TFree(name: String, sort: Sort = Nil) extends Typ {
 | 
| 74121 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 47 |     private lazy val hash: Int = ("TFree", name, sort).hashCode()
 | 
| 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 48 | override def hashCode(): Int = hash | 
| 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 49 | |
| 70846 | 50 | override def toString: String = | 
| 77368 | 51 |       "TFree(" + name + if_proper(sort, "," + sort) + ")"
 | 
| 70846 | 52 | } | 
| 75393 | 53 |   case class TVar(name: Indexname, sort: Sort = Nil) extends Typ {
 | 
| 74121 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 54 |     private lazy val hash: Int = ("TVar", name, sort).hashCode()
 | 
| 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 55 | override def hashCode(): Int = hash | 
| 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 56 | |
| 70846 | 57 | override def toString: String = | 
| 77368 | 58 |       "TVar(" + name + if_proper(sort, "," + sort) + ")"
 | 
| 70846 | 59 | } | 
| 71601 | 60 |   val dummyT: Type = Type("dummy")
 | 
| 43730 | 61 | |
| 75393 | 62 |   sealed abstract class Term {
 | 
| 71221 | 63 | def head: Term = | 
| 64 |       this match {
 | |
| 65 | case App(fun, _) => fun.head | |
| 66 | case _ => this | |
| 67 | } | |
| 68 | } | |
| 75393 | 69 |   case class Const(name: String, typargs: List[Typ] = Nil) extends Term {
 | 
| 74121 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 70 |     private lazy val hash: Int = ("Const", name, typargs).hashCode()
 | 
| 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 71 | override def hashCode(): Int = hash | 
| 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 72 | |
| 70846 | 73 | override def toString: String = | 
| 70897 | 74 | if (this == dummy) "_" | 
| 77368 | 75 |       else "Const(" + name + if_proper(typargs, "," + typargs) + ")"
 | 
| 70846 | 76 | } | 
| 75393 | 77 |   case class Free(name: String, typ: Typ = dummyT) extends Term {
 | 
| 74121 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 78 |     private lazy val hash: Int = ("Free", name, typ).hashCode()
 | 
| 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 79 | override def hashCode(): Int = hash | 
| 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 80 | |
| 70846 | 81 | override def toString: String = | 
| 82 |       "Free(" + name + (if (typ == dummyT) "" else "," + typ) + ")"
 | |
| 83 | } | |
| 75393 | 84 |   case class Var(name: Indexname, typ: Typ = dummyT) extends Term {
 | 
| 74121 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 85 |     private lazy val hash: Int = ("Var", name, typ).hashCode()
 | 
| 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 86 | override def hashCode(): Int = hash | 
| 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 87 | |
| 70846 | 88 | override def toString: String = | 
| 89 |       "Var(" + name + (if (typ == dummyT) "" else "," + typ) + ")"
 | |
| 90 | } | |
| 43730 | 91 | case class Bound(index: Int) extends Term | 
| 75393 | 92 |   case class Abs(name: String, typ: Typ, body: Term) extends Term {
 | 
| 74121 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 93 |     private lazy val hash: Int = ("Abs", name, typ, body).hashCode()
 | 
| 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 94 | override def hashCode(): Int = hash | 
| 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 95 | } | 
| 75393 | 96 |   case class App(fun: Term, arg: Term) extends Term {
 | 
| 74121 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 97 |     private lazy val hash: Int = ("App", fun, arg).hashCode()
 | 
| 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 98 | override def hashCode(): Int = hash | 
| 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 99 | |
| 70850 | 100 | override def toString: String = | 
| 101 |       this match {
 | |
| 102 |         case OFCLASS(ty, c) => "OFCLASS(" + ty + "," + c + ")"
 | |
| 103 |         case _ => "App(" + fun + "," + arg + ")"
 | |
| 104 | } | |
| 105 | } | |
| 68265 | 106 | |
| 70897 | 107 |   def dummy_pattern(ty: Typ): Term = Const("Pure.dummy_pattern", List(ty))
 | 
| 108 | val dummy: Term = dummy_pattern(dummyT) | |
| 109 | ||
| 70533 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 110 | sealed abstract class Proof | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 111 | case object MinProof extends Proof | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 112 | case class PBound(index: Int) extends Proof | 
| 75393 | 113 |   case class Abst(name: String, typ: Typ, body: Proof) extends Proof {
 | 
| 74121 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 114 |     private lazy val hash: Int = ("Abst", name, typ, body).hashCode()
 | 
| 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 115 | override def hashCode(): Int = hash | 
| 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 116 | } | 
| 75393 | 117 |   case class AbsP(name: String, hyp: Term, body: Proof) extends Proof {
 | 
| 74121 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 118 |     private lazy val hash: Int = ("AbsP", name, hyp, body).hashCode()
 | 
| 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 119 | override def hashCode(): Int = hash | 
| 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 120 | } | 
| 75393 | 121 |   case class Appt(fun: Proof, arg: Term) extends Proof {
 | 
| 74121 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 122 |     private lazy val hash: Int = ("Appt", fun, arg).hashCode()
 | 
| 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 123 | override def hashCode(): Int = hash | 
| 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 124 | } | 
| 75393 | 125 |   case class AppP(fun: Proof, arg: Proof) extends Proof {
 | 
| 74121 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 126 |     private lazy val hash: Int = ("AppP", fun, arg).hashCode()
 | 
| 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 127 | override def hashCode(): Int = hash | 
| 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 128 | } | 
| 75393 | 129 |   case class Hyp(hyp: Term) extends Proof {
 | 
| 74121 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 130 |     private lazy val hash: Int = ("Hyp", hyp).hashCode()
 | 
| 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 131 | override def hashCode(): Int = hash | 
| 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 132 | } | 
| 75393 | 133 |   case class PAxm(name: String, types: List[Typ]) extends Proof {
 | 
| 74121 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 134 |     private lazy val hash: Int = ("PAxm", name, types).hashCode()
 | 
| 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 135 | override def hashCode(): Int = hash | 
| 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 136 | } | 
| 75393 | 137 |   case class PClass(typ: Typ, cls: Class) extends Proof {
 | 
| 74121 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 138 |     private lazy val hash: Int = ("PClass", typ, cls).hashCode()
 | 
| 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 139 | override def hashCode(): Int = hash | 
| 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 140 | } | 
| 75393 | 141 |   case class Oracle(name: String, prop: Term, types: List[Typ]) extends Proof {
 | 
| 74121 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 142 |     private lazy val hash: Int = ("Oracle", name, prop, types).hashCode()
 | 
| 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 143 | override def hashCode(): Int = hash | 
| 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 144 | } | 
| 75393 | 145 |   case class PThm(serial: Long, theory_name: String, name: String, types: List[Typ]) extends Proof {
 | 
| 74121 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 146 |     private lazy val hash: Int = ("PThm", serial, theory_name, name, types).hashCode()
 | 
| 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 147 | override def hashCode(): Int = hash | 
| 
bc03b0b82fe6
prefer persistent hash code for cachable items (see also 72b13af7f266);
 wenzelm parents: 
73031diff
changeset | 148 | } | 
| 70533 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 149 | |
| 68265 | 150 | |
| 70850 | 151 | /* type classes within the logic */ | 
| 70385 
68d2c533db9c
more operations: support type classes within the logic;
 wenzelm parents: 
70383diff
changeset | 152 | |
| 75393 | 153 |   object Class_Const {
 | 
| 70850 | 154 | val suffix = "_class" | 
| 155 | def apply(c: Class): String = c + suffix | |
| 156 | def unapply(s: String): Option[Class] = | |
| 157 | if (s.endsWith(suffix)) Some(s.substring(0, s.length - suffix.length)) else None | |
| 158 | } | |
| 70385 
68d2c533db9c
more operations: support type classes within the logic;
 wenzelm parents: 
70383diff
changeset | 159 | |
| 75393 | 160 |   object OFCLASS {
 | 
| 70850 | 161 | def apply(ty: Typ, s: Sort): List[Term] = s.map(c => apply(ty, c)) | 
| 162 | ||
| 163 | def apply(ty: Typ, c: Class): Term = | |
| 164 | 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: 
70383diff
changeset | 165 | |
| 70850 | 166 | def unapply(t: Term): Option[(Typ, String)] = | 
| 167 |       t match {
 | |
| 168 | case App(Const(Class_Const(c), List(ty)), Const(Pure_Thy.TYPE, List(ty1))) | |
| 169 | if ty == ty1 => Some((ty, c)) | |
| 170 | case _ => None | |
| 171 | } | |
| 70385 
68d2c533db9c
more operations: support type classes within the logic;
 wenzelm parents: 
70383diff
changeset | 172 | } | 
| 
68d2c533db9c
more operations: support type classes within the logic;
 wenzelm parents: 
70383diff
changeset | 173 | |
| 
68d2c533db9c
more operations: support type classes within the logic;
 wenzelm parents: 
70383diff
changeset | 174 | |
| 68265 | 175 | |
| 176 | /** cache **/ | |
| 177 | ||
| 75393 | 178 |   object Cache {
 | 
| 73024 | 179 | def make( | 
| 76351 
2cee31cd92f0
generic support for XZ and Zstd compression in Isabelle/Scala;
 wenzelm parents: 
75393diff
changeset | 180 | compress: Compress.Cache = Compress.Cache.make(), | 
| 73026 
237bd6318cc1
more uniform default --- hardly relevant in practice;
 wenzelm parents: 
73024diff
changeset | 181 | max_string: Int = isabelle.Cache.default_max_string, | 
| 73024 | 182 | initial_size: Int = isabelle.Cache.default_initial_size): Cache = | 
| 76351 
2cee31cd92f0
generic support for XZ and Zstd compression in Isabelle/Scala;
 wenzelm parents: 
75393diff
changeset | 183 | new Cache(compress, initial_size, max_string) | 
| 68265 | 184 | |
| 73024 | 185 | val none: Cache = make(max_string = 0) | 
| 186 | } | |
| 187 | ||
| 76351 
2cee31cd92f0
generic support for XZ and Zstd compression in Isabelle/Scala;
 wenzelm parents: 
75393diff
changeset | 188 | class Cache(compress: Compress.Cache, max_string: Int, initial_size: Int) | 
| 
2cee31cd92f0
generic support for XZ and Zstd compression in Isabelle/Scala;
 wenzelm parents: 
75393diff
changeset | 189 |   extends XML.Cache(compress, max_string, initial_size) {
 | 
| 68265 | 190 | protected def cache_indexname(x: Indexname): Indexname = | 
| 70536 | 191 | lookup(x) getOrElse store(Indexname(cache_string(x.name), x.index)) | 
| 68265 | 192 | |
| 193 | protected def cache_sort(x: Sort): Sort = | |
| 73024 | 194 | if (x.isEmpty) Nil | 
| 195 | else lookup(x) getOrElse store(x.map(cache_string)) | |
| 68265 | 196 | |
| 75393 | 197 |     protected def cache_typ(x: Typ): Typ = {
 | 
| 68265 | 198 | if (x == dummyT) dummyT | 
| 199 | else | |
| 200 |         lookup(x) match {
 | |
| 201 | case Some(y) => y | |
| 202 | case None => | |
| 203 |             x match {
 | |
| 70533 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 204 | case Type(name, args) => store(Type(cache_string(name), cache_typs(args))) | 
| 68265 | 205 | case TFree(name, sort) => store(TFree(cache_string(name), cache_sort(sort))) | 
| 206 | case TVar(name, sort) => store(TVar(cache_indexname(name), cache_sort(sort))) | |
| 207 | } | |
| 208 | } | |
| 209 | } | |
| 210 | ||
| 75393 | 211 |     protected def cache_typs(x: List[Typ]): List[Typ] = {
 | 
| 70533 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 212 | if (x.isEmpty) Nil | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 213 |       else {
 | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 214 |         lookup(x) match {
 | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 215 | case Some(y) => y | 
| 71601 | 216 | case None => store(x.map(cache_typ)) | 
| 70533 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 217 | } | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 218 | } | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 219 | } | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 220 | |
| 75393 | 221 |     protected def cache_term(x: Term): Term = {
 | 
| 68265 | 222 |       lookup(x) match {
 | 
| 223 | case Some(y) => y | |
| 224 | case None => | |
| 225 |           x match {
 | |
| 70784 
799437173553
Term_XML.Encode/Decode.term uses Const "typargs";
 wenzelm parents: 
70559diff
changeset | 226 | case Const(name, typargs) => store(Const(cache_string(name), cache_typs(typargs))) | 
| 68265 | 227 | case Free(name, typ) => store(Free(cache_string(name), cache_typ(typ))) | 
| 228 | case Var(name, typ) => store(Var(cache_indexname(name), cache_typ(typ))) | |
| 70536 | 229 | case Bound(_) => store(x) | 
| 68265 | 230 | case Abs(name, typ, body) => | 
| 231 | store(Abs(cache_string(name), cache_typ(typ), cache_term(body))) | |
| 232 | case App(fun, arg) => store(App(cache_term(fun), cache_term(arg))) | |
| 233 | } | |
| 234 | } | |
| 235 | } | |
| 236 | ||
| 75393 | 237 |     protected def cache_proof(x: Proof): Proof = {
 | 
| 70533 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 238 | if (x == MinProof) MinProof | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 239 |       else {
 | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 240 |         lookup(x) match {
 | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 241 | case Some(y) => y | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 242 | case None => | 
| 71781 | 243 |             (x: @unchecked) match {
 | 
| 70536 | 244 | case PBound(_) => store(x) | 
| 70533 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 245 | case Abst(name, typ, body) => | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 246 | store(Abst(cache_string(name), cache_typ(typ), cache_proof(body))) | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 247 | case AbsP(name, hyp, body) => | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 248 | store(AbsP(cache_string(name), cache_term(hyp), cache_proof(body))) | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 249 | case Appt(fun, arg) => store(Appt(cache_proof(fun), cache_term(arg))) | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 250 | case AppP(fun, arg) => store(AppP(cache_proof(fun), cache_proof(arg))) | 
| 70535 | 251 | case Hyp(hyp) => store(Hyp(cache_term(hyp))) | 
| 252 | 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: 
71601diff
changeset | 253 | case PClass(typ, cls) => store(PClass(cache_typ(typ), cache_string(cls))) | 
| 70533 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 254 | case Oracle(name, prop, types) => | 
| 70834 | 255 | store(Oracle(cache_string(name), cache_term(prop), cache_typs(types))) | 
| 70538 
fc9ba6fe367f
clarified PThm: theory_name simplifies retrieval from exports;
 wenzelm parents: 
70537diff
changeset | 256 | case PThm(serial, theory_name, name, types) => | 
| 
fc9ba6fe367f
clarified PThm: theory_name simplifies retrieval from exports;
 wenzelm parents: 
70537diff
changeset | 257 | 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: 
70385diff
changeset | 258 | } | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 259 | } | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 260 | } | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 261 | } | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 262 | |
| 68265 | 263 | // main methods | 
| 73024 | 264 | def indexname(x: Indexname): Indexname = | 
| 265 |       if (no_cache) x else synchronized { cache_indexname(x) }
 | |
| 266 | def sort(x: Sort): Sort = | |
| 267 |       if (no_cache) x else synchronized { cache_sort(x) }
 | |
| 268 | def typ(x: Typ): Typ = | |
| 269 |       if (no_cache) x else synchronized { cache_typ(x) }
 | |
| 270 | def term(x: Term): Term = | |
| 271 |       if (no_cache) x else synchronized { cache_term(x) }
 | |
| 272 | def proof(x: Proof): Proof = | |
| 273 |       if (no_cache) x else synchronized { cache_proof(x) }
 | |
| 68265 | 274 | |
| 275 | def position(x: Position.T): Position.T = | |
| 73024 | 276 | if (no_cache) x | 
| 277 |       else synchronized { x.map({ case (a, b) => (cache_string(a), cache_string(b)) }) }
 | |
| 68265 | 278 | } | 
| 43731 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43730diff
changeset | 279 | } |