| author | wenzelm | 
| Tue, 22 Oct 2019 20:55:13 +0200 | |
| changeset 70923 | 98d9b78b7f47 | 
| parent 70916 | 4c15217d6266 | 
| child 71221 | 4dfb7c937126 | 
| 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 | ||
| 12 | object Term | |
| 13 | {
 | |
| 70383 
38ac2e714729
more operations (avoid clones in Isabelle/MMT and Isabelle/Dedukti);
 wenzelm parents: 
68265diff
changeset | 14 | /* types and terms */ | 
| 
38ac2e714729
more operations (avoid clones in Isabelle/MMT and Isabelle/Dedukti);
 wenzelm parents: 
68265diff
changeset | 15 | |
| 70852 | 16 | sealed case class Indexname(name: String, index: Int = 0) | 
| 70537 
17160e0a60b6
Indexname.toString according to string_of_vname' in ML;
 wenzelm parents: 
70536diff
changeset | 17 |   {
 | 
| 
17160e0a60b6
Indexname.toString according to string_of_vname' in ML;
 wenzelm parents: 
70536diff
changeset | 18 | override def toString: String = | 
| 
17160e0a60b6
Indexname.toString according to string_of_vname' in ML;
 wenzelm parents: 
70536diff
changeset | 19 | if (index == -1) name | 
| 
17160e0a60b6
Indexname.toString according to string_of_vname' in ML;
 wenzelm parents: 
70536diff
changeset | 20 |       else {
 | 
| 
17160e0a60b6
Indexname.toString according to string_of_vname' in ML;
 wenzelm parents: 
70536diff
changeset | 21 | val dot = | 
| 
17160e0a60b6
Indexname.toString according to string_of_vname' in ML;
 wenzelm parents: 
70536diff
changeset | 22 |           Symbol.explode(name).reverse match {
 | 
| 
17160e0a60b6
Indexname.toString according to string_of_vname' in ML;
 wenzelm parents: 
70536diff
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: 
70536diff
changeset | 24 | case c :: _ => Symbol.is_digit(c) | 
| 
17160e0a60b6
Indexname.toString according to string_of_vname' in ML;
 wenzelm parents: 
70536diff
changeset | 25 | case _ => true | 
| 
17160e0a60b6
Indexname.toString according to string_of_vname' in ML;
 wenzelm parents: 
70536diff
changeset | 26 | } | 
| 
17160e0a60b6
Indexname.toString according to string_of_vname' in ML;
 wenzelm parents: 
70536diff
changeset | 27 | if (dot) "?" + name + "." + index | 
| 
17160e0a60b6
Indexname.toString according to string_of_vname' in ML;
 wenzelm parents: 
70536diff
changeset | 28 | else if (index != 0) "?" + name + index | 
| 
17160e0a60b6
Indexname.toString according to string_of_vname' in ML;
 wenzelm parents: 
70536diff
changeset | 29 | else "?" + name | 
| 
17160e0a60b6
Indexname.toString according to string_of_vname' in ML;
 wenzelm parents: 
70536diff
changeset | 30 | } | 
| 
17160e0a60b6
Indexname.toString according to string_of_vname' in ML;
 wenzelm parents: 
70536diff
changeset | 31 | } | 
| 43730 | 32 | |
| 70850 | 33 | type Class = String | 
| 34 | type Sort = List[Class] | |
| 43730 | 35 | |
| 36 | sealed abstract class Typ | |
| 37 | case class Type(name: String, args: List[Typ] = Nil) extends Typ | |
| 70846 | 38 |   {
 | 
| 39 | override def toString: String = | |
| 40 | if (this == dummyT) "_" | |
| 41 |       else "Type(" + name + (if (args.isEmpty) "" else "," + args) + ")"
 | |
| 42 | } | |
| 70851 | 43 | case class TFree(name: String, sort: Sort = Nil) extends Typ | 
| 70846 | 44 |   {
 | 
| 45 | override def toString: String = | |
| 70851 | 46 |       "TFree(" + name + (if (sort.isEmpty) "" else "," + sort) + ")"
 | 
| 70846 | 47 | } | 
| 70851 | 48 | case class TVar(name: Indexname, sort: Sort = Nil) extends Typ | 
| 70846 | 49 |   {
 | 
| 50 | override def toString: String = | |
| 70851 | 51 |       "TVar(" + name + (if (sort.isEmpty) "" else "," + sort) + ")"
 | 
| 70846 | 52 | } | 
| 43730 | 53 |   val dummyT = Type("dummy")
 | 
| 54 | ||
| 55 | sealed abstract class Term | |
| 70846 | 56 | case class Const(name: String, typargs: List[Typ] = Nil) extends Term | 
| 57 |   {
 | |
| 58 | override def toString: String = | |
| 70897 | 59 | if (this == dummy) "_" | 
| 60 |       else "Const(" + name + (if (typargs.isEmpty) "" else "," + typargs) + ")"
 | |
| 70846 | 61 | } | 
| 43730 | 62 | case class Free(name: String, typ: Typ = dummyT) extends Term | 
| 70846 | 63 |   {
 | 
| 64 | override def toString: String = | |
| 65 |       "Free(" + name + (if (typ == dummyT) "" else "," + typ) + ")"
 | |
| 66 | } | |
| 43730 | 67 | case class Var(name: Indexname, typ: Typ = dummyT) extends Term | 
| 70846 | 68 |   {
 | 
| 69 | override def toString: String = | |
| 70 |       "Var(" + name + (if (typ == dummyT) "" else "," + typ) + ")"
 | |
| 71 | } | |
| 43730 | 72 | case class Bound(index: Int) extends Term | 
| 70846 | 73 | case class Abs(name: String, typ: Typ, body: Term) extends Term | 
| 43730 | 74 | case class App(fun: Term, arg: Term) extends Term | 
| 70850 | 75 |   {
 | 
| 76 | override def toString: String = | |
| 77 |       this match {
 | |
| 78 |         case OFCLASS(ty, c) => "OFCLASS(" + ty + "," + c + ")"
 | |
| 79 |         case _ => "App(" + fun + "," + arg + ")"
 | |
| 80 | } | |
| 81 | } | |
| 68265 | 82 | |
| 70897 | 83 |   def dummy_pattern(ty: Typ): Term = Const("Pure.dummy_pattern", List(ty))
 | 
| 84 | val dummy: Term = dummy_pattern(dummyT) | |
| 85 | ||
| 70533 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 86 | sealed abstract class Proof | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 87 | case object MinProof extends Proof | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 88 | case class PBound(index: Int) extends Proof | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 89 | case class Abst(name: String, typ: Typ, body: Proof) extends Proof | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 90 | case class AbsP(name: String, hyp: Term, body: Proof) extends Proof | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 91 | case class Appt(fun: Proof, arg: Term) extends Proof | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 92 | case class AppP(fun: Proof, arg: Proof) extends Proof | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 93 | case class Hyp(hyp: Term) extends Proof | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 94 | case class PAxm(name: String, types: List[Typ]) extends Proof | 
| 70850 | 95 | case class OfClass(typ: Typ, cls: Class) extends Proof | 
| 70834 | 96 | 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: 
70897diff
changeset | 97 | 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: 
70385diff
changeset | 98 | |
| 68265 | 99 | |
| 70850 | 100 | /* type classes within the logic */ | 
| 70385 
68d2c533db9c
more operations: support type classes within the logic;
 wenzelm parents: 
70383diff
changeset | 101 | |
| 70850 | 102 | object Class_Const | 
| 103 |   {
 | |
| 104 | val suffix = "_class" | |
| 105 | def apply(c: Class): String = c + suffix | |
| 106 | def unapply(s: String): Option[Class] = | |
| 107 | if (s.endsWith(suffix)) Some(s.substring(0, s.length - suffix.length)) else None | |
| 108 | } | |
| 70385 
68d2c533db9c
more operations: support type classes within the logic;
 wenzelm parents: 
70383diff
changeset | 109 | |
| 70850 | 110 | object OFCLASS | 
| 111 |   {
 | |
| 112 | def apply(ty: Typ, s: Sort): List[Term] = s.map(c => apply(ty, c)) | |
| 113 | ||
| 114 | def apply(ty: Typ, c: Class): Term = | |
| 115 | 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 | 116 | |
| 70850 | 117 | def unapply(t: Term): Option[(Typ, String)] = | 
| 118 |       t match {
 | |
| 119 | case App(Const(Class_Const(c), List(ty)), Const(Pure_Thy.TYPE, List(ty1))) | |
| 120 | if ty == ty1 => Some((ty, c)) | |
| 121 | case _ => None | |
| 122 | } | |
| 70385 
68d2c533db9c
more operations: support type classes within the logic;
 wenzelm parents: 
70383diff
changeset | 123 | } | 
| 
68d2c533db9c
more operations: support type classes within the logic;
 wenzelm parents: 
70383diff
changeset | 124 | |
| 
68d2c533db9c
more operations: support type classes within the logic;
 wenzelm parents: 
70383diff
changeset | 125 | |
| 68265 | 126 | |
| 127 | /** cache **/ | |
| 128 | ||
| 129 | def make_cache(initial_size: Int = 131071, max_string: Int = Integer.MAX_VALUE): Cache = | |
| 130 | new Cache(initial_size, max_string) | |
| 131 | ||
| 132 | class Cache private[Term](initial_size: Int, max_string: Int) | |
| 133 | extends isabelle.Cache(initial_size, max_string) | |
| 134 |   {
 | |
| 135 | protected def cache_indexname(x: Indexname): Indexname = | |
| 70536 | 136 | lookup(x) getOrElse store(Indexname(cache_string(x.name), x.index)) | 
| 68265 | 137 | |
| 138 | protected def cache_sort(x: Sort): Sort = | |
| 70851 | 139 | if (x.isEmpty) Nil else lookup(x) getOrElse store(x.map(cache_string(_))) | 
| 68265 | 140 | |
| 141 | protected def cache_typ(x: Typ): Typ = | |
| 142 |     {
 | |
| 143 | if (x == dummyT) dummyT | |
| 144 | else | |
| 145 |         lookup(x) match {
 | |
| 146 | case Some(y) => y | |
| 147 | case None => | |
| 148 |             x match {
 | |
| 70533 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 149 | case Type(name, args) => store(Type(cache_string(name), cache_typs(args))) | 
| 68265 | 150 | case TFree(name, sort) => store(TFree(cache_string(name), cache_sort(sort))) | 
| 151 | case TVar(name, sort) => store(TVar(cache_indexname(name), cache_sort(sort))) | |
| 152 | } | |
| 153 | } | |
| 154 | } | |
| 155 | ||
| 70533 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 156 | protected def cache_typs(x: List[Typ]): List[Typ] = | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 157 |     {
 | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 158 | if (x.isEmpty) Nil | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 159 |       else {
 | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 160 |         lookup(x) match {
 | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 161 | case Some(y) => y | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 162 | case None => store(x.map(cache_typ(_))) | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 163 | } | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 164 | } | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 165 | } | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 166 | |
| 68265 | 167 | protected def cache_term(x: Term): Term = | 
| 168 |     {
 | |
| 169 |       lookup(x) match {
 | |
| 170 | case Some(y) => y | |
| 171 | case None => | |
| 172 |           x match {
 | |
| 70784 
799437173553
Term_XML.Encode/Decode.term uses Const "typargs";
 wenzelm parents: 
70559diff
changeset | 173 | case Const(name, typargs) => store(Const(cache_string(name), cache_typs(typargs))) | 
| 68265 | 174 | case Free(name, typ) => store(Free(cache_string(name), cache_typ(typ))) | 
| 175 | case Var(name, typ) => store(Var(cache_indexname(name), cache_typ(typ))) | |
| 70536 | 176 | case Bound(_) => store(x) | 
| 68265 | 177 | case Abs(name, typ, body) => | 
| 178 | store(Abs(cache_string(name), cache_typ(typ), cache_term(body))) | |
| 179 | case App(fun, arg) => store(App(cache_term(fun), cache_term(arg))) | |
| 180 | } | |
| 181 | } | |
| 182 | } | |
| 183 | ||
| 70533 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 184 | protected def cache_proof(x: Proof): Proof = | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 185 |     {
 | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 186 | if (x == MinProof) MinProof | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 187 |       else {
 | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 188 |         lookup(x) match {
 | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 189 | case Some(y) => y | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 190 | case None => | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 191 |             x match {
 | 
| 70536 | 192 | case PBound(_) => store(x) | 
| 70533 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 193 | case Abst(name, typ, body) => | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 194 | store(Abst(cache_string(name), cache_typ(typ), cache_proof(body))) | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 195 | case AbsP(name, hyp, body) => | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 196 | store(AbsP(cache_string(name), cache_term(hyp), cache_proof(body))) | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 197 | 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 | 198 | case AppP(fun, arg) => store(AppP(cache_proof(fun), cache_proof(arg))) | 
| 70535 | 199 | case Hyp(hyp) => store(Hyp(cache_term(hyp))) | 
| 200 | case PAxm(name, types) => store(PAxm(cache_string(name), cache_typs(types))) | |
| 70533 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 201 | case OfClass(typ, cls) => store(OfClass(cache_typ(typ), cache_string(cls))) | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 202 | case Oracle(name, prop, types) => | 
| 70834 | 203 | 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 | 204 | case PThm(serial, theory_name, name, types) => | 
| 
fc9ba6fe367f
clarified PThm: theory_name simplifies retrieval from exports;
 wenzelm parents: 
70537diff
changeset | 205 | 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 | 206 | } | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 207 | } | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 208 | } | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 209 | } | 
| 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 210 | |
| 68265 | 211 | // main methods | 
| 212 |     def indexname(x: Indexname): Indexname = synchronized { cache_indexname(x) }
 | |
| 213 |     def sort(x: Sort): Sort = synchronized { cache_sort(x) }
 | |
| 214 |     def typ(x: Typ): Typ = synchronized { cache_typ(x) }
 | |
| 215 |     def term(x: Term): Term = synchronized { cache_term(x) }
 | |
| 70533 
031620901fcd
support for (fully reconstructed) proof terms in Scala;
 wenzelm parents: 
70385diff
changeset | 216 |     def proof(x: Proof): Proof = synchronized { cache_proof(x) }
 | 
| 68265 | 217 | |
| 218 | def position(x: Position.T): Position.T = | |
| 219 |       synchronized { x.map({ case (a, b) => (cache_string(a), cache_string(b)) }) }
 | |
| 220 | } | |
| 43731 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43730diff
changeset | 221 | } |