| author | wenzelm | 
| Mon, 01 Apr 2024 15:37:55 +0200 | |
| changeset 80064 | 0d94dd2fd2d0 | 
| parent 77368 | 7c57d9586f4c | 
| permissions | -rw-r--r-- | 
| 36947 | 1 | /* Title: Pure/Isar/keyword.scala | 
| 29449 
6e7745d35a30
added outer_keyword.scala: Isar command keyword classification;
 wenzelm parents: diff
changeset | 2 | Author: Makarius | 
| 
6e7745d35a30
added outer_keyword.scala: Isar command keyword classification;
 wenzelm parents: diff
changeset | 3 | |
| 58900 | 4 | Isar keyword classification. | 
| 29449 
6e7745d35a30
added outer_keyword.scala: Isar command keyword classification;
 wenzelm parents: diff
changeset | 5 | */ | 
| 
6e7745d35a30
added outer_keyword.scala: Isar command keyword classification;
 wenzelm parents: diff
changeset | 6 | |
| 
6e7745d35a30
added outer_keyword.scala: Isar command keyword classification;
 wenzelm parents: diff
changeset | 7 | package isabelle | 
| 
6e7745d35a30
added outer_keyword.scala: Isar command keyword classification;
 wenzelm parents: diff
changeset | 8 | |
| 
6e7745d35a30
added outer_keyword.scala: Isar command keyword classification;
 wenzelm parents: diff
changeset | 9 | |
| 75393 | 10 | object Keyword {
 | 
| 58900 | 11 | /** keyword classification **/ | 
| 12 | ||
| 36681 | 13 | /* kinds */ | 
| 14 | ||
| 29449 
6e7745d35a30
added outer_keyword.scala: Isar command keyword classification;
 wenzelm parents: diff
changeset | 15 | val DIAG = "diag" | 
| 58999 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 wenzelm parents: 
58928diff
changeset | 16 | val DOCUMENT_HEADING = "document_heading" | 
| 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 wenzelm parents: 
58928diff
changeset | 17 | val DOCUMENT_BODY = "document_body" | 
| 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 wenzelm parents: 
58928diff
changeset | 18 | val DOCUMENT_RAW = "document_raw" | 
| 46967 
499d9bbd8de9
uniform keyword names within ML/Scala -- produce elisp names via external conversion;
 wenzelm parents: 
46123diff
changeset | 19 | val THY_BEGIN = "thy_begin" | 
| 
499d9bbd8de9
uniform keyword names within ML/Scala -- produce elisp names via external conversion;
 wenzelm parents: 
46123diff
changeset | 20 | val THY_END = "thy_end" | 
| 
499d9bbd8de9
uniform keyword names within ML/Scala -- produce elisp names via external conversion;
 wenzelm parents: 
46123diff
changeset | 21 | val THY_DECL = "thy_decl" | 
| 58800 
bfed1c26caed
explicit keyword category for commands that may start a block;
 wenzelm parents: 
57837diff
changeset | 22 | val THY_DECL_BLOCK = "thy_decl_block" | 
| 69913 | 23 | val THY_DEFN = "thy_defn" | 
| 24 | val THY_STMT = "thy_stmt" | |
| 48867 
e9beabf045ab
some support for inlining file content into outer syntax token language;
 wenzelm parents: 
46969diff
changeset | 25 | val THY_LOAD = "thy_load" | 
| 46967 
499d9bbd8de9
uniform keyword names within ML/Scala -- produce elisp names via external conversion;
 wenzelm parents: 
46123diff
changeset | 26 | val THY_GOAL = "thy_goal" | 
| 69913 | 27 | val THY_GOAL_DEFN = "thy_goal_defn" | 
| 28 | val THY_GOAL_STMT = "thy_goal_stmt" | |
| 29449 
6e7745d35a30
added outer_keyword.scala: Isar command keyword classification;
 wenzelm parents: diff
changeset | 29 | val QED = "qed" | 
| 53571 
e58ca0311c0f
more explicit indication of 'done' as proof script element;
 wenzelm parents: 
53371diff
changeset | 30 | val QED_SCRIPT = "qed_script" | 
| 46967 
499d9bbd8de9
uniform keyword names within ML/Scala -- produce elisp names via external conversion;
 wenzelm parents: 
46123diff
changeset | 31 | val QED_BLOCK = "qed_block" | 
| 
499d9bbd8de9
uniform keyword names within ML/Scala -- produce elisp names via external conversion;
 wenzelm parents: 
46123diff
changeset | 32 | val QED_GLOBAL = "qed_global" | 
| 
499d9bbd8de9
uniform keyword names within ML/Scala -- produce elisp names via external conversion;
 wenzelm parents: 
46123diff
changeset | 33 | val PRF_GOAL = "prf_goal" | 
| 
499d9bbd8de9
uniform keyword names within ML/Scala -- produce elisp names via external conversion;
 wenzelm parents: 
46123diff
changeset | 34 | val PRF_BLOCK = "prf_block" | 
| 60694 
b3fa4a8cdb5f
clarified text folds: proof ... qed counts as extra block;
 wenzelm parents: 
60692diff
changeset | 35 | val NEXT_BLOCK = "next_block" | 
| 46967 
499d9bbd8de9
uniform keyword names within ML/Scala -- produce elisp names via external conversion;
 wenzelm parents: 
46123diff
changeset | 36 | val PRF_OPEN = "prf_open" | 
| 
499d9bbd8de9
uniform keyword names within ML/Scala -- produce elisp names via external conversion;
 wenzelm parents: 
46123diff
changeset | 37 | val PRF_CLOSE = "prf_close" | 
| 
499d9bbd8de9
uniform keyword names within ML/Scala -- produce elisp names via external conversion;
 wenzelm parents: 
46123diff
changeset | 38 | val PRF_CHAIN = "prf_chain" | 
| 
499d9bbd8de9
uniform keyword names within ML/Scala -- produce elisp names via external conversion;
 wenzelm parents: 
46123diff
changeset | 39 | val PRF_DECL = "prf_decl" | 
| 
499d9bbd8de9
uniform keyword names within ML/Scala -- produce elisp names via external conversion;
 wenzelm parents: 
46123diff
changeset | 40 | val PRF_ASM = "prf_asm" | 
| 
499d9bbd8de9
uniform keyword names within ML/Scala -- produce elisp names via external conversion;
 wenzelm parents: 
46123diff
changeset | 41 | val PRF_ASM_GOAL = "prf_asm_goal" | 
| 
499d9bbd8de9
uniform keyword names within ML/Scala -- produce elisp names via external conversion;
 wenzelm parents: 
46123diff
changeset | 42 | val PRF_SCRIPT = "prf_script" | 
| 60624 | 43 | val PRF_SCRIPT_GOAL = "prf_script_goal" | 
| 44 | val PRF_SCRIPT_ASM_GOAL = "prf_script_asm_goal" | |
| 63441 | 45 | |
| 46 | val BEFORE_COMMAND = "before_command" | |
| 63430 | 47 | val QUASI_COMMAND = "quasi_command" | 
| 29449 
6e7745d35a30
added outer_keyword.scala: Isar command keyword classification;
 wenzelm parents: diff
changeset | 48 | |
| 36681 | 49 | |
| 59700 | 50 | /* command categories */ | 
| 36681 | 51 | |
| 70123 | 52 | val vacuous = Set(DIAG, DOCUMENT_HEADING, DOCUMENT_BODY, DOCUMENT_RAW) | 
| 58999 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 wenzelm parents: 
58928diff
changeset | 53 | |
| 57835 | 54 | val diag = Set(DIAG) | 
| 55 | ||
| 58999 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 wenzelm parents: 
58928diff
changeset | 56 | val document_heading = Set(DOCUMENT_HEADING) | 
| 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 wenzelm parents: 
58928diff
changeset | 57 | val document_body = Set(DOCUMENT_BODY) | 
| 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 wenzelm parents: 
58928diff
changeset | 58 | val document_raw = Set(DOCUMENT_RAW) | 
| 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 wenzelm parents: 
58928diff
changeset | 59 | val document = Set(DOCUMENT_HEADING, DOCUMENT_BODY, DOCUMENT_RAW) | 
| 57835 | 60 | |
| 59700 | 61 | val theory_begin = Set(THY_BEGIN) | 
| 62 | val theory_end = Set(THY_END) | |
| 63 | ||
| 59701 | 64 | val theory_load = Set(THY_LOAD) | 
| 65 | ||
| 69913 | 66 | val theory = | 
| 67 | Set(THY_BEGIN, THY_END, THY_LOAD, THY_DECL, THY_DECL_BLOCK, THY_DEFN, THY_STMT, | |
| 68 | THY_GOAL, THY_GOAL_DEFN, THY_GOAL_STMT) | |
| 57837 | 69 | |
| 58800 
bfed1c26caed
explicit keyword category for commands that may start a block;
 wenzelm parents: 
57837diff
changeset | 70 | val theory_block = Set(THY_BEGIN, THY_DECL_BLOCK) | 
| 
bfed1c26caed
explicit keyword category for commands that may start a block;
 wenzelm parents: 
57837diff
changeset | 71 | |
| 69913 | 72 | val theory_body = | 
| 73 | Set(THY_LOAD, THY_DECL, THY_DECL_BLOCK, THY_DEFN, THY_STMT, | |
| 74 | THY_GOAL, THY_GOAL_DEFN, THY_GOAL_STMT) | |
| 57835 | 75 | |
| 69917 | 76 | val theory_defn = Set(THY_DEFN, THY_GOAL_DEFN) | 
| 77 | ||
| 63479 | 78 | val prf_script = Set(PRF_SCRIPT) | 
| 79 | ||
| 40456 | 80 | val proof = | 
| 60694 
b3fa4a8cdb5f
clarified text folds: proof ... qed counts as extra block;
 wenzelm parents: 
60692diff
changeset | 81 | Set(QED, QED_SCRIPT, QED_BLOCK, QED_GLOBAL, PRF_GOAL, PRF_BLOCK, NEXT_BLOCK, PRF_OPEN, | 
| 
b3fa4a8cdb5f
clarified text folds: proof ... qed counts as extra block;
 wenzelm parents: 
60692diff
changeset | 82 | PRF_CLOSE, PRF_CHAIN, PRF_DECL, PRF_ASM, PRF_ASM_GOAL, PRF_SCRIPT, PRF_SCRIPT_GOAL, | 
| 60624 | 83 | PRF_SCRIPT_ASM_GOAL) | 
| 57835 | 84 | |
| 85 | val proof_body = | |
| 60694 
b3fa4a8cdb5f
clarified text folds: proof ... qed counts as extra block;
 wenzelm parents: 
60692diff
changeset | 86 | Set(DIAG, DOCUMENT_HEADING, DOCUMENT_BODY, DOCUMENT_RAW, PRF_BLOCK, NEXT_BLOCK, PRF_OPEN, | 
| 
b3fa4a8cdb5f
clarified text folds: proof ... qed counts as extra block;
 wenzelm parents: 
60692diff
changeset | 87 | PRF_CLOSE, PRF_CHAIN, PRF_DECL, PRF_ASM, PRF_ASM_GOAL, PRF_SCRIPT, PRF_SCRIPT_GOAL, | 
| 60624 | 88 | PRF_SCRIPT_ASM_GOAL) | 
| 57835 | 89 | |
| 72669 | 90 | val proof_asm = Set(PRF_ASM, PRF_ASM_GOAL) | 
| 91 | val improper = Set(QED_SCRIPT, PRF_SCRIPT, PRF_SCRIPT_GOAL, PRF_SCRIPT_ASM_GOAL) | |
| 92 | ||
| 69913 | 93 | val theory_goal = Set(THY_GOAL, THY_GOAL_DEFN, THY_GOAL_STMT) | 
| 60624 | 94 | val proof_goal = Set(PRF_GOAL, PRF_ASM_GOAL, PRF_SCRIPT_GOAL, PRF_SCRIPT_ASM_GOAL) | 
| 57835 | 95 | val qed = Set(QED, QED_SCRIPT, QED_BLOCK) | 
| 96 | val qed_global = Set(QED_GLOBAL) | |
| 58900 | 97 | |
| 60692 | 98 | val proof_open = proof_goal + PRF_OPEN | 
| 99 | val proof_close = qed + PRF_CLOSE | |
| 63428 
005b490f0ce2
indentation in reminiscence to Proof General (see proof-indent.el);
 wenzelm parents: 
63424diff
changeset | 100 | val proof_enclose = Set(PRF_BLOCK, NEXT_BLOCK, QED_BLOCK, PRF_CLOSE) | 
| 60692 | 101 | |
| 63603 | 102 | val close_structure = Set(NEXT_BLOCK, QED_BLOCK, PRF_CLOSE, THY_END) | 
| 103 | ||
| 58900 | 104 | |
| 105 | ||
| 106 | /** keyword tables **/ | |
| 107 | ||
| 72764 | 108 | sealed case class Spec( | 
| 109 | kind: String = "", | |
| 72765 | 110 | kind_pos: Position.T = Position.none, | 
| 72764 | 111 | load_command: String = "", | 
| 72765 | 112 | load_command_pos: Position.T = Position.none, | 
| 75393 | 113 | tags: List[String] = Nil | 
| 114 |   ) {
 | |
| 65385 | 115 | override def toString: String = | 
| 116 | kind + | |
| 77368 | 117 |         if_proper(load_command, " (" + quote(load_command) + ")") +
 | 
| 118 |         if_proper(tags, tags.map(quote).mkString(" % ", " % ", ""))
 | |
| 72764 | 119 | |
| 120 | def map(f: String => String): Spec = | |
| 121 | copy(kind = f(kind), load_command = f(load_command), tags = tags.map(f)) | |
| 65384 | 122 | } | 
| 63429 | 123 | |
| 75393 | 124 |   object Keywords {
 | 
| 58900 | 125 | def empty: Keywords = new Keywords() | 
| 126 | } | |
| 127 | ||
| 128 | class Keywords private( | |
| 63441 | 129 | val kinds: Map[String, String] = Map.empty, | 
| 75393 | 130 | val load_commands: Map[String, String] = Map.empty | 
| 131 |   ) {
 | |
| 132 |     override def toString: String = {
 | |
| 63441 | 133 | val entries = | 
| 134 |         for ((name, kind) <- kinds.toList.sortBy(_._1)) yield {
 | |
| 72748 
04d5f6d769a7
more flexible syntax for theory load commands via Isabelle/Scala;
 wenzelm parents: 
72740diff
changeset | 135 | val load_decl = | 
| 
04d5f6d769a7
more flexible syntax for theory load commands via Isabelle/Scala;
 wenzelm parents: 
72740diff
changeset | 136 |             load_commands.get(name) match {
 | 
| 
04d5f6d769a7
more flexible syntax for theory load commands via Isabelle/Scala;
 wenzelm parents: 
72740diff
changeset | 137 |               case Some(load_command) => " (" + quote(load_command) + ")"
 | 
| 
04d5f6d769a7
more flexible syntax for theory load commands via Isabelle/Scala;
 wenzelm parents: 
72740diff
changeset | 138 | case None => "" | 
| 
04d5f6d769a7
more flexible syntax for theory load commands via Isabelle/Scala;
 wenzelm parents: 
72740diff
changeset | 139 | } | 
| 77368 | 140 | val kind_decl = if_proper(kind, " :: " + quote(kind) + load_decl) | 
| 63441 | 141 | quote(name) + kind_decl | 
| 63430 | 142 | } | 
| 63441 | 143 |       entries.mkString("keywords\n  ", " and\n  ", "")
 | 
| 58900 | 144 | } | 
| 145 | ||
| 146 | ||
| 59073 | 147 | /* merge */ | 
| 148 | ||
| 67090 
0ec94bb9cec4
clarified lazy lexicons within Keywords: measurable speedup of Sessions.deps;
 wenzelm parents: 
66919diff
changeset | 149 | def is_empty: Boolean = kinds.isEmpty | 
| 59073 | 150 | |
| 151 | def ++ (other: Keywords): Keywords = | |
| 152 | if (this eq other) this | |
| 153 | else if (is_empty) other | |
| 154 |       else {
 | |
| 63441 | 155 | val kinds1 = | 
| 156 | if (kinds eq other.kinds) kinds | |
| 157 | else if (kinds.isEmpty) other.kinds | |
| 73359 | 158 |           else other.kinds.foldLeft(kinds) { case (m, e) => if (m.isDefinedAt(e._1)) m else m + e }
 | 
| 63441 | 159 | val load_commands1 = | 
| 160 | if (load_commands eq other.load_commands) load_commands | |
| 161 | else if (load_commands.isEmpty) other.load_commands | |
| 73359 | 162 |           else {
 | 
| 163 |             other.load_commands.foldLeft(load_commands) {
 | |
| 164 | case (m, e) => if (m.isDefinedAt(e._1)) m else m + e | |
| 165 | } | |
| 166 | } | |
| 67090 
0ec94bb9cec4
clarified lazy lexicons within Keywords: measurable speedup of Sessions.deps;
 wenzelm parents: 
66919diff
changeset | 167 | new Keywords(kinds1, load_commands1) | 
| 59073 | 168 | } | 
| 169 | ||
| 170 | ||
| 58902 | 171 | /* add keywords */ | 
| 172 | ||
| 75393 | 173 |     def + (name: String, kind: String = "", load_command: String = ""): Keywords = {
 | 
| 63441 | 174 | val kinds1 = kinds + (name -> kind) | 
| 175 | val load_commands1 = | |
| 66919 | 176 |         if (kind == THY_LOAD) {
 | 
| 71601 | 177 | if (!Symbol.iterator(name).forall(Symbol.is_ascii)) | 
| 66919 | 178 |             error("Bad theory load command " + quote(name))
 | 
| 72748 
04d5f6d769a7
more flexible syntax for theory load commands via Isabelle/Scala;
 wenzelm parents: 
72740diff
changeset | 179 | load_commands + (name -> load_command) | 
| 66919 | 180 | } | 
| 63441 | 181 | else load_commands | 
| 67090 
0ec94bb9cec4
clarified lazy lexicons within Keywords: measurable speedup of Sessions.deps;
 wenzelm parents: 
66919diff
changeset | 182 | new Keywords(kinds1, load_commands1) | 
| 63441 | 183 | } | 
| 58902 | 184 | |
| 58928 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58907diff
changeset | 185 | def add_keywords(header: Thy_Header.Keywords): Keywords = | 
| 73359 | 186 |       header.foldLeft(this) {
 | 
| 65384 | 187 | case (keywords, (name, spec)) => | 
| 72764 | 188 | if (spec.kind.isEmpty) | 
| 65384 | 189 | keywords + Symbol.decode(name) + Symbol.encode(name) | 
| 190 | else | |
| 191 | keywords + | |
| 72748 
04d5f6d769a7
more flexible syntax for theory load commands via Isabelle/Scala;
 wenzelm parents: 
72740diff
changeset | 192 | (Symbol.decode(name), spec.kind, spec.load_command) + | 
| 
04d5f6d769a7
more flexible syntax for theory load commands via Isabelle/Scala;
 wenzelm parents: 
72740diff
changeset | 193 | (Symbol.encode(name), spec.kind, spec.load_command) | 
| 58928 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58907diff
changeset | 194 | } | 
| 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58907diff
changeset | 195 | |
| 58902 | 196 | |
| 58901 | 197 | /* command kind */ | 
| 58900 | 198 | |
| 63424 | 199 | def is_command(token: Token, check_kind: String => Boolean): Boolean = | 
| 200 | token.is_command && | |
| 63441 | 201 |         (kinds.get(token.source) match { case Some(k) => check_kind(k) case None => false })
 | 
| 202 | ||
| 203 | def is_before_command(token: Token): Boolean = | |
| 204 | token.is_keyword && kinds.get(token.source) == Some(BEFORE_COMMAND) | |
| 59701 | 205 | |
| 63430 | 206 | def is_quasi_command(token: Token): Boolean = | 
| 63441 | 207 | token.is_keyword && kinds.get(token.source) == Some(QUASI_COMMAND) | 
| 63430 | 208 | |
| 63809 | 209 | def is_indent_command(token: Token): Boolean = | 
| 210 | token.is_begin_or_command || is_quasi_command(token) | |
| 211 | ||
| 58900 | 212 | |
| 67090 
0ec94bb9cec4
clarified lazy lexicons within Keywords: measurable speedup of Sessions.deps;
 wenzelm parents: 
66919diff
changeset | 213 | /* lexicons */ | 
| 
0ec94bb9cec4
clarified lazy lexicons within Keywords: measurable speedup of Sessions.deps;
 wenzelm parents: 
66919diff
changeset | 214 | |
| 
0ec94bb9cec4
clarified lazy lexicons within Keywords: measurable speedup of Sessions.deps;
 wenzelm parents: 
66919diff
changeset | 215 | private def make_lexicon(is_minor: Boolean): Scan.Lexicon = | 
| 73359 | 216 |       kinds.foldLeft(Scan.Lexicon.empty) {
 | 
| 217 | case (lex, (name, kind)) => | |
| 218 | if ((kind == "" || kind == BEFORE_COMMAND || kind == QUASI_COMMAND) == is_minor) | |
| 219 | lex + name | |
| 220 | else lex | |
| 221 | } | |
| 67090 
0ec94bb9cec4
clarified lazy lexicons within Keywords: measurable speedup of Sessions.deps;
 wenzelm parents: 
66919diff
changeset | 222 | |
| 
0ec94bb9cec4
clarified lazy lexicons within Keywords: measurable speedup of Sessions.deps;
 wenzelm parents: 
66919diff
changeset | 223 | lazy val minor: Scan.Lexicon = make_lexicon(true) | 
| 
0ec94bb9cec4
clarified lazy lexicons within Keywords: measurable speedup of Sessions.deps;
 wenzelm parents: 
66919diff
changeset | 224 | lazy val major: Scan.Lexicon = make_lexicon(false) | 
| 58900 | 225 | } | 
| 29449 
6e7745d35a30
added outer_keyword.scala: Isar command keyword classification;
 wenzelm parents: diff
changeset | 226 | } |