| author | paulson <lp15@cam.ac.uk> | 
| Sun, 15 Nov 2020 13:06:24 +0000 | |
| changeset 72608 | ad45ae49be85 | 
| parent 68845 | 3b2daa7bf9f4 | 
| child 72692 | 22aeec526ffd | 
| permissions | -rw-r--r-- | 
| 
57905
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents:  
diff
changeset
 | 
1  | 
/* Title: Pure/PIDE/command_span.scala  | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents:  
diff
changeset
 | 
2  | 
Author: Makarius  | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents:  
diff
changeset
 | 
3  | 
|
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents:  
diff
changeset
 | 
4  | 
Syntactic representation of command spans.  | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents:  
diff
changeset
 | 
5  | 
*/  | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents:  
diff
changeset
 | 
6  | 
|
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents:  
diff
changeset
 | 
7  | 
package isabelle  | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents:  
diff
changeset
 | 
8  | 
|
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents:  
diff
changeset
 | 
9  | 
|
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents:  
diff
changeset
 | 
10  | 
import scala.collection.mutable  | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents:  
diff
changeset
 | 
11  | 
|
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents:  
diff
changeset
 | 
12  | 
|
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents:  
diff
changeset
 | 
13  | 
object Command_Span  | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents:  
diff
changeset
 | 
14  | 
{
 | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents:  
diff
changeset
 | 
15  | 
  sealed abstract class Kind {
 | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents:  
diff
changeset
 | 
16  | 
override def toString: String =  | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents:  
diff
changeset
 | 
17  | 
      this match {
 | 
| 65717 | 18  | 
case Command_Span(name, _) => proper_string(name) getOrElse "<command>"  | 
| 
57905
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents:  
diff
changeset
 | 
19  | 
case Ignored_Span => "<ignored>"  | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents:  
diff
changeset
 | 
20  | 
case Malformed_Span => "<malformed>"  | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents:  
diff
changeset
 | 
21  | 
}  | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents:  
diff
changeset
 | 
22  | 
}  | 
| 57910 | 23  | 
case class Command_Span(name: String, pos: Position.T) extends Kind  | 
| 
57905
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents:  
diff
changeset
 | 
24  | 
case object Ignored_Span extends Kind  | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents:  
diff
changeset
 | 
25  | 
case object Malformed_Span extends Kind  | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents:  
diff
changeset
 | 
26  | 
|
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents:  
diff
changeset
 | 
27  | 
sealed case class Span(kind: Kind, content: List[Token])  | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents:  
diff
changeset
 | 
28  | 
  {
 | 
| 59735 | 29  | 
def name: String =  | 
30  | 
      kind match { case Command_Span(name, _) => name case _ => "" }
 | 
|
| 58802 | 31  | 
|
| 59735 | 32  | 
def position: Position.T =  | 
33  | 
      kind match { case Command_Span(_, pos) => pos case _ => Position.none }
 | 
|
34  | 
||
| 67895 | 35  | 
def keyword_pos(start: Token.Pos): Token.Pos =  | 
36  | 
      kind match {
 | 
|
37  | 
case _: Command_Span =>  | 
|
38  | 
(start /: content.iterator.takeWhile(tok => !tok.is_command))(_.advance(_))  | 
|
39  | 
case _ => start  | 
|
40  | 
}  | 
|
41  | 
||
| 
68845
 
3b2daa7bf9f4
support Thy_Element in Scala, following ML version;
 
wenzelm 
parents: 
68840 
diff
changeset
 | 
42  | 
def is_kind(keywords: Keyword.Keywords, pred: String => Boolean, other: Boolean): Boolean =  | 
| 68840 | 43  | 
      keywords.kinds.get(name) match {
 | 
44  | 
case Some(k) => pred(k)  | 
|
| 
68845
 
3b2daa7bf9f4
support Thy_Element in Scala, following ML version;
 
wenzelm 
parents: 
68840 
diff
changeset
 | 
45  | 
case None => other  | 
| 68840 | 46  | 
}  | 
47  | 
||
| 63606 | 48  | 
def is_begin: Boolean = content.exists(_.is_begin)  | 
49  | 
def is_end: Boolean = content.exists(_.is_end)  | 
|
50  | 
||
| 59735 | 51  | 
def length: Int = (0 /: content)(_ + _.source.length)  | 
| 59705 | 52  | 
|
| 
57905
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents:  
diff
changeset
 | 
53  | 
def compact_source: (String, Span) =  | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents:  
diff
changeset
 | 
54  | 
    {
 | 
| 59735 | 55  | 
val source = Token.implode(content)  | 
| 
57905
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents:  
diff
changeset
 | 
56  | 
val content1 = new mutable.ListBuffer[Token]  | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents:  
diff
changeset
 | 
57  | 
var i = 0  | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents:  
diff
changeset
 | 
58  | 
      for (Token(kind, s) <- content) {
 | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents:  
diff
changeset
 | 
59  | 
val n = s.length  | 
| 59735 | 60  | 
val s1 = source.substring(i, i + n)  | 
| 
57905
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents:  
diff
changeset
 | 
61  | 
content1 += Token(kind, s1)  | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents:  
diff
changeset
 | 
62  | 
i += n  | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents:  
diff
changeset
 | 
63  | 
}  | 
| 59735 | 64  | 
(source, Span(kind, content1.toList))  | 
| 
57905
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents:  
diff
changeset
 | 
65  | 
}  | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents:  
diff
changeset
 | 
66  | 
}  | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents:  
diff
changeset
 | 
67  | 
|
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents:  
diff
changeset
 | 
68  | 
val empty: Span = Span(Ignored_Span, Nil)  | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents:  
diff
changeset
 | 
69  | 
|
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents:  
diff
changeset
 | 
70  | 
def unparsed(source: String): Span =  | 
| 
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents:  
diff
changeset
 | 
71  | 
Span(Malformed_Span, List(Token(Token.Kind.UNPARSED, source)))  | 
| 65335 | 72  | 
|
73  | 
||
74  | 
/* XML data representation */  | 
|
75  | 
||
76  | 
def encode: XML.Encode.T[Span] = (span: Span) =>  | 
|
77  | 
  {
 | 
|
78  | 
import XML.Encode._  | 
|
79  | 
val kind: T[Kind] =  | 
|
80  | 
variant(List(  | 
|
81  | 
        { case Command_Span(a, b) => (List(a), properties(b)) },
 | 
|
82  | 
        { case Ignored_Span => (Nil, Nil) },
 | 
|
83  | 
        { case Malformed_Span => (Nil, Nil) }))
 | 
|
84  | 
pair(kind, list(Token.encode))((span.kind, span.content))  | 
|
85  | 
}  | 
|
86  | 
||
87  | 
def decode: XML.Decode.T[Span] = (body: XML.Body) =>  | 
|
88  | 
  {
 | 
|
89  | 
import XML.Decode._  | 
|
90  | 
val kind: T[Kind] =  | 
|
91  | 
variant(List(  | 
|
92  | 
        { case (List(a), b) => Command_Span(a, properties(b)) },
 | 
|
93  | 
        { case (Nil, Nil) => Ignored_Span },
 | 
|
94  | 
        { case (Nil, Nil) => Malformed_Span }))
 | 
|
95  | 
val (k, toks) = pair(kind, list(Token.decode))(body)  | 
|
96  | 
Span(k, toks)  | 
|
97  | 
}  | 
|
| 
57905
 
c0c5652e796e
separate module Command_Span: mostly syntactic representation;
 
wenzelm 
parents:  
diff
changeset
 | 
98  | 
}  |