| author | bulwahn | 
| Thu, 16 Sep 2010 13:49:12 +0200 | |
| changeset 39466 | f3c5da707f30 | 
| parent 37072 | 9105c8237c7a | 
| child 40533 | e38e80686ce5 | 
| permissions | -rw-r--r-- | 
| 31763 | 1  | 
/* Title: Pure/Thy/completion.scala  | 
2  | 
Author: Makarius  | 
|
3  | 
||
4  | 
Completion of symbols and keywords.  | 
|
5  | 
*/  | 
|
6  | 
||
7  | 
package isabelle  | 
|
8  | 
||
9  | 
import scala.util.matching.Regex  | 
|
10  | 
import scala.util.parsing.combinator.RegexParsers  | 
|
11  | 
||
12  | 
||
13  | 
private object Completion  | 
|
14  | 
{
 | 
|
15  | 
/** word completion **/  | 
|
16  | 
||
17  | 
val word_regex = "[a-zA-Z0-9_']+".r  | 
|
18  | 
def is_word(s: CharSequence): Boolean = word_regex.pattern.matcher(s).matches  | 
|
19  | 
||
20  | 
object Parse extends RegexParsers  | 
|
21  | 
  {
 | 
|
22  | 
override val whiteSpace = "".r  | 
|
23  | 
||
| 
37072
 
9105c8237c7a
renamed "rev" to "reverse" following usual Scala conventions;
 
wenzelm 
parents: 
36012 
diff
changeset
 | 
24  | 
def reverse_symbol: Parser[String] = """>[A-Za-z0-9_']+<\\""".r  | 
| 
 
9105c8237c7a
renamed "rev" to "reverse" following usual Scala conventions;
 
wenzelm 
parents: 
36012 
diff
changeset
 | 
25  | 
    def reverse_symb: Parser[String] = """[A-Za-z0-9_']{2,}<\\""".r
 | 
| 31763 | 26  | 
    def word: Parser[String] = "[a-zA-Z0-9_']{3,}".r
 | 
27  | 
||
28  | 
def read(in: CharSequence): Option[String] =  | 
|
29  | 
    {
 | 
|
| 
37072
 
9105c8237c7a
renamed "rev" to "reverse" following usual Scala conventions;
 
wenzelm 
parents: 
36012 
diff
changeset
 | 
30  | 
val reverse_in = new Library.Reverse(in)  | 
| 
 
9105c8237c7a
renamed "rev" to "reverse" following usual Scala conventions;
 
wenzelm 
parents: 
36012 
diff
changeset
 | 
31  | 
      parse((reverse_symbol | reverse_symb | word) ^^ (_.reverse), reverse_in) match {
 | 
| 31763 | 32  | 
case Success(result, _) => Some(result)  | 
33  | 
case _ => None  | 
|
34  | 
}  | 
|
35  | 
}  | 
|
36  | 
}  | 
|
37  | 
}  | 
|
38  | 
||
39  | 
class Completion  | 
|
40  | 
{
 | 
|
41  | 
/* representation */  | 
|
42  | 
||
| 31780 | 43  | 
protected val words_lex = Scan.Lexicon()  | 
44  | 
protected val words_map = Map[String, String]()  | 
|
| 31763 | 45  | 
|
| 31780 | 46  | 
protected val abbrevs_lex = Scan.Lexicon()  | 
47  | 
protected val abbrevs_map = Map[String, (String, String)]()  | 
|
| 31763 | 48  | 
|
49  | 
||
50  | 
/* adding stuff */  | 
|
51  | 
||
52  | 
def + (keyword: String): Completion =  | 
|
53  | 
  {
 | 
|
54  | 
val old = this  | 
|
55  | 
    new Completion {
 | 
|
56  | 
override val words_lex = old.words_lex + keyword  | 
|
57  | 
override val words_map = old.words_map + (keyword -> keyword)  | 
|
58  | 
override val abbrevs_lex = old.abbrevs_lex  | 
|
59  | 
override val abbrevs_map = old.abbrevs_map  | 
|
60  | 
}  | 
|
61  | 
}  | 
|
62  | 
||
63  | 
def + (symbols: Symbol.Interpretation): Completion =  | 
|
64  | 
  {
 | 
|
65  | 
val words =  | 
|
66  | 
(for ((x, _) <- symbols.names) yield (x, x)).toList :::  | 
|
67  | 
(for ((x, y) <- symbols.names) yield (y, x)).toList :::  | 
|
68  | 
(for ((x, y) <- symbols.abbrevs if Completion.is_word(y)) yield (y, x)).toList  | 
|
69  | 
||
70  | 
val abbrs =  | 
|
71  | 
(for ((x, y) <- symbols.abbrevs if !Completion.is_word(y))  | 
|
| 31765 | 72  | 
yield (y.reverse.toString, (y, x))).toList  | 
| 31763 | 73  | 
|
74  | 
val old = this  | 
|
75  | 
    new Completion {
 | 
|
76  | 
override val words_lex = old.words_lex ++ words.map(_._1)  | 
|
| 31765 | 77  | 
override val words_map = old.words_map ++ words  | 
| 31763 | 78  | 
override val abbrevs_lex = old.abbrevs_lex ++ abbrs.map(_._1)  | 
79  | 
override val abbrevs_map = old.abbrevs_map ++ abbrs  | 
|
80  | 
}  | 
|
81  | 
}  | 
|
82  | 
||
83  | 
||
84  | 
/* complete */  | 
|
85  | 
||
| 31765 | 86  | 
def complete(line: CharSequence): Option[(String, List[String])] =  | 
| 31763 | 87  | 
  {
 | 
| 34141 | 88  | 
    abbrevs_lex.parse(abbrevs_lex.keyword, new Library.Reverse(line)) match {
 | 
| 
37072
 
9105c8237c7a
renamed "rev" to "reverse" following usual Scala conventions;
 
wenzelm 
parents: 
36012 
diff
changeset
 | 
89  | 
case abbrevs_lex.Success(reverse_a, _) =>  | 
| 
 
9105c8237c7a
renamed "rev" to "reverse" following usual Scala conventions;
 
wenzelm 
parents: 
36012 
diff
changeset
 | 
90  | 
val (word, c) = abbrevs_map(reverse_a)  | 
| 
35004
 
b89a31957950
filter out identical completions only after symbols.decode -- recover completion of literal symbols (e.g. \<AA>);
 
wenzelm 
parents: 
34141 
diff
changeset
 | 
91  | 
Some(word, List(c))  | 
| 31765 | 92  | 
case _ =>  | 
93  | 
        Completion.Parse.read(line) match {
 | 
|
94  | 
case Some(word) =>  | 
|
| 
35004
 
b89a31957950
filter out identical completions only after symbols.decode -- recover completion of literal symbols (e.g. \<AA>);
 
wenzelm 
parents: 
34141 
diff
changeset
 | 
95  | 
            words_lex.completions(word).map(words_map(_)) match {
 | 
| 31765 | 96  | 
case Nil => None  | 
| 36012 | 97  | 
case cs => Some(word, cs.sortWith(_ < _))  | 
| 31765 | 98  | 
}  | 
99  | 
case None => None  | 
|
100  | 
}  | 
|
101  | 
}  | 
|
| 31763 | 102  | 
}  | 
103  | 
}  |