| author | Christian Sternagel | 
| Wed, 29 Aug 2012 10:35:05 +0900 | |
| changeset 49078 | 398e8fddabb0 | 
| parent 46712 | 8650d9a95736 | 
| child 53251 | 7facc08da806 | 
| 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  | 
||
| 46624 | 13  | 
object Completion  | 
| 31763 | 14  | 
{
 | 
| 46625 | 15  | 
val empty: Completion = new Completion()  | 
| 46624 | 16  | 
def init(): Completion = empty.add_symbols()  | 
17  | 
||
18  | 
||
| 31763 | 19  | 
/** word completion **/  | 
20  | 
||
| 46624 | 21  | 
private val word_regex = "[a-zA-Z0-9_']+".r  | 
22  | 
private def is_word(s: CharSequence): Boolean = word_regex.pattern.matcher(s).matches  | 
|
| 31763 | 23  | 
|
| 46624 | 24  | 
private object Parse extends RegexParsers  | 
| 31763 | 25  | 
  {
 | 
26  | 
override val whiteSpace = "".r  | 
|
27  | 
||
| 43462 | 28  | 
def reverse_symbol: Parser[String] = """>[A-Za-z0-9_']+\^?<\\""".r  | 
29  | 
    def reverse_symb: Parser[String] = """[A-Za-z0-9_']{2,}\^?<\\""".r
 | 
|
| 
40533
 
e38e80686ce5
somewhat adhoc replacement for 'thus' and 'hence';
 
wenzelm 
parents: 
37072 
diff
changeset
 | 
30  | 
    def word: Parser[String] = "[a-zA-Z0-9_']{2,}".r
 | 
| 31763 | 31  | 
|
32  | 
def read(in: CharSequence): Option[String] =  | 
|
33  | 
    {
 | 
|
| 
37072
 
9105c8237c7a
renamed "rev" to "reverse" following usual Scala conventions;
 
wenzelm 
parents: 
36012 
diff
changeset
 | 
34  | 
val reverse_in = new Library.Reverse(in)  | 
| 
 
9105c8237c7a
renamed "rev" to "reverse" following usual Scala conventions;
 
wenzelm 
parents: 
36012 
diff
changeset
 | 
35  | 
      parse((reverse_symbol | reverse_symb | word) ^^ (_.reverse), reverse_in) match {
 | 
| 31763 | 36  | 
case Success(result, _) => Some(result)  | 
37  | 
case _ => None  | 
|
38  | 
}  | 
|
39  | 
}  | 
|
40  | 
}  | 
|
41  | 
}  | 
|
42  | 
||
| 46712 | 43  | 
final class Completion private(  | 
| 46625 | 44  | 
words_lex: Scan.Lexicon = Scan.Lexicon.empty,  | 
45  | 
words_map: Map[String, String] = Map.empty,  | 
|
46  | 
abbrevs_lex: Scan.Lexicon = Scan.Lexicon.empty,  | 
|
47  | 
abbrevs_map: Map[String, (String, String)] = Map.empty)  | 
|
| 31763 | 48  | 
{
 | 
49  | 
/* adding stuff */  | 
|
50  | 
||
| 
40533
 
e38e80686ce5
somewhat adhoc replacement for 'thus' and 'hence';
 
wenzelm 
parents: 
37072 
diff
changeset
 | 
51  | 
def + (keyword: String, replace: String): Completion =  | 
| 46624 | 52  | 
new Completion(  | 
53  | 
words_lex + keyword,  | 
|
54  | 
words_map + (keyword -> replace),  | 
|
55  | 
abbrevs_lex,  | 
|
56  | 
abbrevs_map)  | 
|
| 31763 | 57  | 
|
| 
40533
 
e38e80686ce5
somewhat adhoc replacement for 'thus' and 'hence';
 
wenzelm 
parents: 
37072 
diff
changeset
 | 
58  | 
def + (keyword: String): Completion = this + (keyword, keyword)  | 
| 
 
e38e80686ce5
somewhat adhoc replacement for 'thus' and 'hence';
 
wenzelm 
parents: 
37072 
diff
changeset
 | 
59  | 
|
| 46624 | 60  | 
private def add_symbols(): Completion =  | 
| 31763 | 61  | 
  {
 | 
62  | 
val words =  | 
|
| 
43695
 
5130dfe1b7be
simplified Symbol based on lazy Symbol.Interpretation -- reduced odd "functorial style";
 
wenzelm 
parents: 
43462 
diff
changeset
 | 
63  | 
(for ((x, _) <- Symbol.names) yield (x, x)).toList :::  | 
| 
 
5130dfe1b7be
simplified Symbol based on lazy Symbol.Interpretation -- reduced odd "functorial style";
 
wenzelm 
parents: 
43462 
diff
changeset
 | 
64  | 
(for ((x, y) <- Symbol.names) yield (y, x)).toList :::  | 
| 
 
5130dfe1b7be
simplified Symbol based on lazy Symbol.Interpretation -- reduced odd "functorial style";
 
wenzelm 
parents: 
43462 
diff
changeset
 | 
65  | 
(for ((x, y) <- Symbol.abbrevs if Completion.is_word(y)) yield (y, x)).toList  | 
| 31763 | 66  | 
|
67  | 
val abbrs =  | 
|
| 
43695
 
5130dfe1b7be
simplified Symbol based on lazy Symbol.Interpretation -- reduced odd "functorial style";
 
wenzelm 
parents: 
43462 
diff
changeset
 | 
68  | 
(for ((x, y) <- Symbol.abbrevs if !Completion.is_word(y))  | 
| 31765 | 69  | 
yield (y.reverse.toString, (y, x))).toList  | 
| 31763 | 70  | 
|
| 46624 | 71  | 
new Completion(  | 
72  | 
words_lex ++ words.map(_._1),  | 
|
73  | 
words_map ++ words,  | 
|
74  | 
abbrevs_lex ++ abbrs.map(_._1),  | 
|
75  | 
abbrevs_map ++ abbrs)  | 
|
| 31763 | 76  | 
}  | 
77  | 
||
78  | 
||
79  | 
/* complete */  | 
|
80  | 
||
| 31765 | 81  | 
def complete(line: CharSequence): Option[(String, List[String])] =  | 
| 31763 | 82  | 
  {
 | 
| 34141 | 83  | 
    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
 | 
84  | 
case abbrevs_lex.Success(reverse_a, _) =>  | 
| 
 
9105c8237c7a
renamed "rev" to "reverse" following usual Scala conventions;
 
wenzelm 
parents: 
36012 
diff
changeset
 | 
85  | 
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
 | 
86  | 
Some(word, List(c))  | 
| 31765 | 87  | 
case _ =>  | 
88  | 
        Completion.Parse.read(line) match {
 | 
|
89  | 
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
 | 
90  | 
            words_lex.completions(word).map(words_map(_)) match {
 | 
| 31765 | 91  | 
case Nil => None  | 
| 45900 | 92  | 
case cs => Some(word, cs.sorted)  | 
| 31765 | 93  | 
}  | 
94  | 
case None => None  | 
|
95  | 
}  | 
|
96  | 
}  | 
|
| 31763 | 97  | 
}  | 
98  | 
}  |