| author | eberlm | 
| Fri, 17 Jun 2016 11:33:52 +0200 | |
| changeset 63319 | bc8793d7bd21 | 
| parent 63019 | 80ef19b51493 | 
| child 63640 | c273583f0203 | 
| permissions | -rw-r--r-- | 
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
32738 
diff
changeset
 | 
1  | 
(* Title: Pure/Isar/token.ML  | 
| 5825 | 2  | 
Author: Markus Wenzel, TU Muenchen  | 
3  | 
||
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
32738 
diff
changeset
 | 
4  | 
Outer token syntax for Isabelle/Isar.  | 
| 5825 | 5  | 
*)  | 
6  | 
||
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
32738 
diff
changeset
 | 
7  | 
signature TOKEN =  | 
| 5825 | 8  | 
sig  | 
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
32738 
diff
changeset
 | 
9  | 
datatype kind =  | 
| 59081 | 10  | 
(*immediate source*)  | 
11  | 
Command | Keyword | Ident | Long_Ident | Sym_Ident | Var | Type_Ident | Type_Var | Nat |  | 
|
12  | 
Float | Space |  | 
|
13  | 
(*delimited content*)  | 
|
14  | 
String | Alt_String | Verbatim | Cartouche | Comment |  | 
|
15  | 
(*special content*)  | 
|
| 61819 | 16  | 
Error of string | EOF  | 
| 58012 | 17  | 
val str_of_kind: kind -> string  | 
| 
55788
 
67699e08e969
store blobs / inlined files as separate text lines: smaller values are more healthy for the Poly/ML RTS and allow implicit sharing;
 
wenzelm 
parents: 
55750 
diff
changeset
 | 
18  | 
  type file = {src_path: Path.T, lines: string list, digest: SHA1.digest, pos: Position.T}
 | 
| 58012 | 19  | 
type T  | 
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
20  | 
type src = T list  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
21  | 
  type name_value = {name: string, kind: string, print: Proof.context -> Markup.T * xstring}
 | 
| 27814 | 22  | 
datatype value =  | 
| 58012 | 23  | 
Source of src |  | 
| 
57944
 
fff8d328da56
more informative Token.Name with history of morphisms;
 
wenzelm 
parents: 
57942 
diff
changeset
 | 
24  | 
Literal of bool * Markup.T |  | 
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
25  | 
Name of name_value * morphism |  | 
| 
57944
 
fff8d328da56
more informative Token.Name with history of morphisms;
 
wenzelm 
parents: 
57942 
diff
changeset
 | 
26  | 
Typ of typ |  | 
| 
 
fff8d328da56
more informative Token.Name with history of morphisms;
 
wenzelm 
parents: 
57942 
diff
changeset
 | 
27  | 
Term of term |  | 
| 
57942
 
e5bec882fdd0
more informative Token.Fact: retain name of dynamic fact (without selection);
 
wenzelm 
parents: 
56202 
diff
changeset
 | 
28  | 
Fact of string option * thm list |  | 
| 
57944
 
fff8d328da56
more informative Token.Name with history of morphisms;
 
wenzelm 
parents: 
57942 
diff
changeset
 | 
29  | 
Attribute of morphism -> attribute |  | 
| 58017 | 30  | 
Declaration of declaration |  | 
| 
57944
 
fff8d328da56
more informative Token.Name with history of morphisms;
 
wenzelm 
parents: 
57942 
diff
changeset
 | 
31  | 
Files of file Exn.result list  | 
| 55708 | 32  | 
val pos_of: T -> Position.T  | 
| 
55709
 
4e5a83a46ded
clarified Token.range_of in accordance to Symbol_Pos.range;
 
wenzelm 
parents: 
55708 
diff
changeset
 | 
33  | 
val range_of: T list -> Position.range  | 
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
32738 
diff
changeset
 | 
34  | 
val eof: T  | 
| 
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
32738 
diff
changeset
 | 
35  | 
val is_eof: T -> bool  | 
| 
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
32738 
diff
changeset
 | 
36  | 
val not_eof: T -> bool  | 
| 
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
32738 
diff
changeset
 | 
37  | 
val stopper: T Scan.stopper  | 
| 
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
32738 
diff
changeset
 | 
38  | 
val kind_of: T -> kind  | 
| 
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
32738 
diff
changeset
 | 
39  | 
val is_kind: kind -> T -> bool  | 
| 
46811
 
03a2dc9e0624
clarified command span: include trailing whitespace/comments and thus reduce number of ignored spans with associated transactions and states (factor 2);
 
wenzelm 
parents: 
45666 
diff
changeset
 | 
40  | 
val is_command: T -> bool  | 
| 
59924
 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 
wenzelm 
parents: 
59913 
diff
changeset
 | 
41  | 
val keyword_with: (string -> bool) -> T -> bool  | 
| 
59939
 
7d46aa03696e
support for 'restricted' modifier: only qualified accesses outside the local scope;
 
wenzelm 
parents: 
59924 
diff
changeset
 | 
42  | 
val is_command_modifier: T -> bool  | 
| 
59924
 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 
wenzelm 
parents: 
59913 
diff
changeset
 | 
43  | 
val ident_with: (string -> bool) -> T -> bool  | 
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
32738 
diff
changeset
 | 
44  | 
val is_proper: T -> bool  | 
| 
51266
 
3007d0bc9cb1
unified Command.is_proper in ML with Scala (see also 123be08eed88);
 
wenzelm 
parents: 
50239 
diff
changeset
 | 
45  | 
val is_improper: T -> bool  | 
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
32738 
diff
changeset
 | 
46  | 
val is_comment: T -> bool  | 
| 
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
32738 
diff
changeset
 | 
47  | 
val is_begin_ignore: T -> bool  | 
| 
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
32738 
diff
changeset
 | 
48  | 
val is_end_ignore: T -> bool  | 
| 
48749
 
c197b3c3e7fa
some attempts to keep malformed syntax errors focussed, without too much red spilled onto the document view;
 
wenzelm 
parents: 
48743 
diff
changeset
 | 
49  | 
val is_error: T -> bool  | 
| 
48771
 
2ea997196d04
clarified Command.range vs. Command.proper_range according to Scala version, which is potentially relevant for command status markup;
 
wenzelm 
parents: 
48764 
diff
changeset
 | 
50  | 
val is_space: T -> bool  | 
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
32738 
diff
changeset
 | 
51  | 
val is_blank: T -> bool  | 
| 
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
32738 
diff
changeset
 | 
52  | 
val is_newline: T -> bool  | 
| 59795 | 53  | 
val content_of: T -> string  | 
| 59809 | 54  | 
val input_of: T -> Input.source  | 
| 59795 | 55  | 
val inner_syntax_of: T -> string  | 
| 56202 | 56  | 
val keyword_markup: bool * Markup.T -> string -> Markup.T  | 
| 
55915
 
607948c90bf0
suppress short abbreviations more uniformly, for outer and quasi-outer syntax;
 
wenzelm 
parents: 
55914 
diff
changeset
 | 
57  | 
val completion_report: T -> Position.report_text list  | 
| 59125 | 58  | 
val reports: Keyword.keywords -> T -> Position.report_text list  | 
59  | 
val markups: Keyword.keywords -> T -> Markup.T list  | 
|
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
32738 
diff
changeset
 | 
60  | 
val unparse: T -> string  | 
| 55745 | 61  | 
val print: T -> string  | 
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
32738 
diff
changeset
 | 
62  | 
val text_of: T -> string * string  | 
| 
54520
 
cee77d2e9582
release file errors at runtime: Command.eval instead of Command.read;
 
wenzelm 
parents: 
54519 
diff
changeset
 | 
63  | 
val get_files: T -> file Exn.result list  | 
| 
 
cee77d2e9582
release file errors at runtime: Command.eval instead of Command.read;
 
wenzelm 
parents: 
54519 
diff
changeset
 | 
64  | 
val put_files: file Exn.result list -> T -> T  | 
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
32738 
diff
changeset
 | 
65  | 
val get_value: T -> value option  | 
| 61822 | 66  | 
val reports_of_value: T -> Position.report list  | 
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
67  | 
val name_value: name_value -> value  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
68  | 
val get_name: T -> name_value option  | 
| 59646 | 69  | 
val declare_maxidx: T -> Proof.context -> Proof.context  | 
| 61820 | 70  | 
val map_facts: (string option -> thm list -> thm list) -> T -> T  | 
| 
58011
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57944 
diff
changeset
 | 
71  | 
val transform: morphism -> T -> T  | 
| 
55914
 
c5b752d549e3
clarified init_assignable: make double-sure that initial values are reset;
 
wenzelm 
parents: 
55828 
diff
changeset
 | 
72  | 
val init_assignable: T -> T  | 
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
73  | 
val assign: value option -> T -> T  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
74  | 
  val evaluate: ('a -> value) -> (T -> 'a) -> T -> 'a
 | 
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
32738 
diff
changeset
 | 
75  | 
val closure: T -> T  | 
| 58012 | 76  | 
val pretty_value: Proof.context -> T -> Pretty.T  | 
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
77  | 
val name_of_src: src -> string * Position.T  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
78  | 
val args_of_src: src -> T list  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
79  | 
val checked_src: src -> bool  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
80  | 
val check_src: Proof.context -> (Proof.context -> 'a Name_Space.table) -> src -> src * 'a  | 
| 
58011
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57944 
diff
changeset
 | 
81  | 
val pretty_src: Proof.context -> src -> Pretty.T  | 
| 27814 | 82  | 
val ident_or_symbolic: string -> bool  | 
| 61614 | 83  | 
val source': bool -> Keyword.keywords -> (Symbol_Pos.T, 'a) Source.source ->  | 
84  | 
(T, (Symbol_Pos.T, 'a) Source.source) Source.source  | 
|
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
32738 
diff
changeset
 | 
85  | 
val source_proper: (T, 'a) Source.source -> (T, (T, 'a) Source.source) Source.source  | 
| 
58904
 
f49c4f883c58
eliminated pointless dynamic keywords (TTY legacy);
 
wenzelm 
parents: 
58903 
diff
changeset
 | 
86  | 
val source: Keyword.keywords ->  | 
| 58864 | 87  | 
Position.T -> (Symbol.symbol, 'a) Source.source -> (T,  | 
88  | 
(Symbol_Pos.T, Position.T * (Symbol.symbol, 'a) Source.source) Source.source) Source.source  | 
|
| 
58904
 
f49c4f883c58
eliminated pointless dynamic keywords (TTY legacy);
 
wenzelm 
parents: 
58903 
diff
changeset
 | 
89  | 
val source_strict: Keyword.keywords ->  | 
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
32738 
diff
changeset
 | 
90  | 
Position.T -> (Symbol.symbol, 'a) Source.source -> (T,  | 
| 30573 | 91  | 
(Symbol_Pos.T, Position.T * (Symbol.symbol, 'a) Source.source) Source.source) Source.source  | 
| 61471 | 92  | 
val read_cartouche: Symbol_Pos.T list -> T  | 
| 59083 | 93  | 
val explode: Keyword.keywords -> Position.T -> string -> T list  | 
| 
59085
 
08a6901eb035
clarified define_command: send tokens more directly, without requiring keywords in ML;
 
wenzelm 
parents: 
59083 
diff
changeset
 | 
94  | 
val make: (int * int) * string -> Position.T -> T * Position.T  | 
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
95  | 
val make_string: string * Position.T -> T  | 
| 63019 | 96  | 
val make_int: int -> T list  | 
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
97  | 
val make_src: string * Position.T -> T list -> src  | 
| 
58011
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57944 
diff
changeset
 | 
98  | 
type 'a parser = T list -> 'a * T list  | 
| 
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57944 
diff
changeset
 | 
99  | 
type 'a context_parser = Context.generic * T list -> 'a * (Context.generic * T list)  | 
| 58903 | 100  | 
val read_no_commands: Keyword.keywords -> 'a parser -> Symbol_Pos.T list -> 'a list  | 
101  | 
val read_antiq: Keyword.keywords -> 'a parser -> Symbol_Pos.T list * Position.T -> 'a  | 
|
| 
58011
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57944 
diff
changeset
 | 
102  | 
val syntax_generic: 'a context_parser -> src -> Context.generic -> 'a * Context.generic  | 
| 
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57944 
diff
changeset
 | 
103  | 
val syntax: 'a context_parser -> src -> Proof.context -> 'a * Proof.context  | 
| 5825 | 104  | 
end;  | 
105  | 
||
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
32738 
diff
changeset
 | 
106  | 
structure Token: TOKEN =  | 
| 5825 | 107  | 
struct  | 
108  | 
||
109  | 
(** tokens **)  | 
|
110  | 
||
| 58012 | 111  | 
(* token kind *)  | 
| 5825 | 112  | 
|
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
32738 
diff
changeset
 | 
113  | 
datatype kind =  | 
| 59081 | 114  | 
(*immediate source*)  | 
115  | 
Command | Keyword | Ident | Long_Ident | Sym_Ident | Var | Type_Ident | Type_Var | Nat |  | 
|
116  | 
Float | Space |  | 
|
117  | 
(*delimited content*)  | 
|
118  | 
String | Alt_String | Verbatim | Cartouche | Comment |  | 
|
119  | 
(*special content*)  | 
|
| 61819 | 120  | 
Error of string | EOF;  | 
| 5825 | 121  | 
|
122  | 
val str_of_kind =  | 
|
| 7026 | 123  | 
fn Command => "command"  | 
124  | 
| Keyword => "keyword"  | 
|
| 5825 | 125  | 
| Ident => "identifier"  | 
| 59081 | 126  | 
| Long_Ident => "long identifier"  | 
127  | 
| Sym_Ident => "symbolic identifier"  | 
|
| 5825 | 128  | 
| Var => "schematic variable"  | 
| 59081 | 129  | 
| Type_Ident => "type variable"  | 
130  | 
| Type_Var => "schematic type variable"  | 
|
| 
40290
 
47f572aff50a
support for floating-point tokens in outer syntax (coinciding with inner syntax version);
 
wenzelm 
parents: 
38229 
diff
changeset
 | 
131  | 
| Nat => "natural number"  | 
| 
 
47f572aff50a
support for floating-point tokens in outer syntax (coinciding with inner syntax version);
 
wenzelm 
parents: 
38229 
diff
changeset
 | 
132  | 
| Float => "floating-point number"  | 
| 59081 | 133  | 
| Space => "white space"  | 
| 55103 | 134  | 
| String => "quoted string"  | 
| 59081 | 135  | 
| Alt_String => "back-quoted string"  | 
| 5825 | 136  | 
| Verbatim => "verbatim text"  | 
| 55103 | 137  | 
| Cartouche => "text cartouche"  | 
| 
7682
 
46de8064c93c
added Space, Comment token kinds (keep actual text);
 
wenzelm 
parents: 
7477 
diff
changeset
 | 
138  | 
| Comment => "comment text"  | 
| 
23729
 
d1ba656978c5
separated Malformed (symbolic char) from Error (bad input);
 
wenzelm 
parents: 
23721 
diff
changeset
 | 
139  | 
| Error _ => "bad input"  | 
| 
48911
 
5debc3e4fa81
tuned messages: end-of-input rarely means physical end-of-file from the past;
 
wenzelm 
parents: 
48905 
diff
changeset
 | 
140  | 
| EOF => "end-of-input";  | 
| 5825 | 141  | 
|
| 
59085
 
08a6901eb035
clarified define_command: send tokens more directly, without requiring keywords in ML;
 
wenzelm 
parents: 
59083 
diff
changeset
 | 
142  | 
val immediate_kinds =  | 
| 
 
08a6901eb035
clarified define_command: send tokens more directly, without requiring keywords in ML;
 
wenzelm 
parents: 
59083 
diff
changeset
 | 
143  | 
Vector.fromList  | 
| 
 
08a6901eb035
clarified define_command: send tokens more directly, without requiring keywords in ML;
 
wenzelm 
parents: 
59083 
diff
changeset
 | 
144  | 
[Command, Keyword, Ident, Long_Ident, Sym_Ident, Var, Type_Ident, Type_Var, Nat, Float, Space];  | 
| 
 
08a6901eb035
clarified define_command: send tokens more directly, without requiring keywords in ML;
 
wenzelm 
parents: 
59083 
diff
changeset
 | 
145  | 
|
| 59081 | 146  | 
val delimited_kind = member (op =) [String, Alt_String, Verbatim, Cartouche, Comment];  | 
| 
55828
 
42ac3cfb89f6
clarified language markup: added "delimited" property;
 
wenzelm 
parents: 
55788 
diff
changeset
 | 
147  | 
|
| 5825 | 148  | 
|
| 58012 | 149  | 
(* datatype token *)  | 
150  | 
||
151  | 
(*The value slot assigns an (optional) internal value to a token,  | 
|
152  | 
usually as a side-effect of special scanner setup (see also  | 
|
153  | 
args.ML). Note that an assignable ref designates an intermediate  | 
|
154  | 
state of internalization -- it is NOT meant to persist.*)  | 
|
155  | 
||
156  | 
type file = {src_path: Path.T, lines: string list, digest: SHA1.digest, pos: Position.T};
 | 
|
157  | 
||
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
158  | 
type name_value = {name: string, kind: string, print: Proof.context -> Markup.T * xstring};
 | 
| 58012 | 159  | 
|
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
160  | 
datatype T = Token of (Symbol_Pos.text * Position.range) * (kind * string) * slot  | 
| 58012 | 161  | 
|
162  | 
and slot =  | 
|
163  | 
Slot |  | 
|
164  | 
Value of value option |  | 
|
165  | 
Assignable of value option Unsynchronized.ref  | 
|
166  | 
||
167  | 
and value =  | 
|
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
168  | 
Source of T list |  | 
| 58012 | 169  | 
Literal of bool * Markup.T |  | 
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
170  | 
Name of name_value * morphism |  | 
| 58012 | 171  | 
Typ of typ |  | 
172  | 
Term of term |  | 
|
173  | 
Fact of string option * thm list | (*optional name for dynamic fact, i.e. fact "variable"*)  | 
|
174  | 
Attribute of morphism -> attribute |  | 
|
| 58017 | 175  | 
Declaration of declaration |  | 
| 58012 | 176  | 
Files of file Exn.result list;  | 
177  | 
||
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
178  | 
type src = T list;  | 
| 58012 | 179  | 
|
180  | 
||
| 
27733
 
d3d7038fb7b5
abstract type Scan.stopper, position taken from last input token;
 
wenzelm 
parents: 
27663 
diff
changeset
 | 
181  | 
(* position *)  | 
| 5825 | 182  | 
|
| 55708 | 183  | 
fun pos_of (Token ((_, (pos, _)), _, _)) = pos;  | 
184  | 
fun end_pos_of (Token ((_, (_, pos)), _, _)) = pos;  | 
|
| 27663 | 185  | 
|
| 
55709
 
4e5a83a46ded
clarified Token.range_of in accordance to Symbol_Pos.range;
 
wenzelm 
parents: 
55708 
diff
changeset
 | 
186  | 
fun range_of (toks as tok :: _) =  | 
| 
 
4e5a83a46ded
clarified Token.range_of in accordance to Symbol_Pos.range;
 
wenzelm 
parents: 
55708 
diff
changeset
 | 
187  | 
let val pos' = end_pos_of (List.last toks)  | 
| 62797 | 188  | 
in Position.range (pos_of tok, pos') end  | 
| 
55709
 
4e5a83a46ded
clarified Token.range_of in accordance to Symbol_Pos.range;
 
wenzelm 
parents: 
55708 
diff
changeset
 | 
189  | 
| range_of [] = Position.no_range;  | 
| 5825 | 190  | 
|
191  | 
||
| 58855 | 192  | 
(* stopper *)  | 
| 
27733
 
d3d7038fb7b5
abstract type Scan.stopper, position taken from last input token;
 
wenzelm 
parents: 
27663 
diff
changeset
 | 
193  | 
|
| 27814 | 194  | 
fun mk_eof pos = Token (("", (pos, Position.none)), (EOF, ""), Slot);
 | 
| 
27733
 
d3d7038fb7b5
abstract type Scan.stopper, position taken from last input token;
 
wenzelm 
parents: 
27663 
diff
changeset
 | 
195  | 
val eof = mk_eof Position.none;  | 
| 
 
d3d7038fb7b5
abstract type Scan.stopper, position taken from last input token;
 
wenzelm 
parents: 
27663 
diff
changeset
 | 
196  | 
|
| 27814 | 197  | 
fun is_eof (Token (_, (EOF, _), _)) = true  | 
| 
27733
 
d3d7038fb7b5
abstract type Scan.stopper, position taken from last input token;
 
wenzelm 
parents: 
27663 
diff
changeset
 | 
198  | 
| is_eof _ = false;  | 
| 
 
d3d7038fb7b5
abstract type Scan.stopper, position taken from last input token;
 
wenzelm 
parents: 
27663 
diff
changeset
 | 
199  | 
|
| 
 
d3d7038fb7b5
abstract type Scan.stopper, position taken from last input token;
 
wenzelm 
parents: 
27663 
diff
changeset
 | 
200  | 
val not_eof = not o is_eof;  | 
| 
 
d3d7038fb7b5
abstract type Scan.stopper, position taken from last input token;
 
wenzelm 
parents: 
27663 
diff
changeset
 | 
201  | 
|
| 
27752
 
ea7d573e565f
removed obsolete range_of (already included in position);
 
wenzelm 
parents: 
27747 
diff
changeset
 | 
202  | 
val stopper =  | 
| 55708 | 203  | 
Scan.stopper (fn [] => eof | toks => mk_eof (end_pos_of (List.last toks))) is_eof;  | 
| 
27733
 
d3d7038fb7b5
abstract type Scan.stopper, position taken from last input token;
 
wenzelm 
parents: 
27663 
diff
changeset
 | 
204  | 
|
| 
 
d3d7038fb7b5
abstract type Scan.stopper, position taken from last input token;
 
wenzelm 
parents: 
27663 
diff
changeset
 | 
205  | 
|
| 5825 | 206  | 
(* kind of token *)  | 
207  | 
||
| 27814 | 208  | 
fun kind_of (Token (_, (k, _), _)) = k;  | 
209  | 
fun is_kind k (Token (_, (k', _), _)) = k = k';  | 
|
| 5825 | 210  | 
|
| 
46811
 
03a2dc9e0624
clarified command span: include trailing whitespace/comments and thus reduce number of ignored spans with associated transactions and states (factor 2);
 
wenzelm 
parents: 
45666 
diff
changeset
 | 
211  | 
val is_command = is_kind Command;  | 
| 
59123
 
e68e44836d04
imitate command markup and rendering of Isabelle/jEdit in HTML output;
 
wenzelm 
parents: 
59112 
diff
changeset
 | 
212  | 
|
| 27814 | 213  | 
fun keyword_with pred (Token (_, (Keyword, x), _)) = pred x  | 
| 7026 | 214  | 
| keyword_with _ _ = false;  | 
| 5825 | 215  | 
|
| 
59990
 
a81dc82ecba3
clarified keyword 'qualified' in accordance to a similar keyword from Haskell (despite unrelated Binding.qualified in Isabelle/ML);
 
wenzelm 
parents: 
59939 
diff
changeset
 | 
216  | 
val is_command_modifier = keyword_with (fn x => x = "private" orelse x = "qualified");  | 
| 
59924
 
801b979ec0c2
more general notion of command span: command keyword not necessarily at start;
 
wenzelm 
parents: 
59913 
diff
changeset
 | 
217  | 
|
| 27814 | 218  | 
fun ident_with pred (Token (_, (Ident, x), _)) = pred x  | 
| 16029 | 219  | 
| ident_with _ _ = false;  | 
220  | 
||
| 27814 | 221  | 
fun is_proper (Token (_, (Space, _), _)) = false  | 
222  | 
| is_proper (Token (_, (Comment, _), _)) = false  | 
|
| 5825 | 223  | 
| is_proper _ = true;  | 
224  | 
||
| 
51266
 
3007d0bc9cb1
unified Command.is_proper in ML with Scala (see also 123be08eed88);
 
wenzelm 
parents: 
50239 
diff
changeset
 | 
225  | 
val is_improper = not o is_proper;  | 
| 
 
3007d0bc9cb1
unified Command.is_proper in ML with Scala (see also 123be08eed88);
 
wenzelm 
parents: 
50239 
diff
changeset
 | 
226  | 
|
| 27814 | 227  | 
fun is_comment (Token (_, (Comment, _), _)) = true  | 
| 17069 | 228  | 
| is_comment _ = false;  | 
229  | 
||
| 27814 | 230  | 
fun is_begin_ignore (Token (_, (Comment, "<"), _)) = true  | 
| 8580 | 231  | 
| is_begin_ignore _ = false;  | 
232  | 
||
| 27814 | 233  | 
fun is_end_ignore (Token (_, (Comment, ">"), _)) = true  | 
| 8580 | 234  | 
| is_end_ignore _ = false;  | 
235  | 
||
| 
48749
 
c197b3c3e7fa
some attempts to keep malformed syntax errors focussed, without too much red spilled onto the document view;
 
wenzelm 
parents: 
48743 
diff
changeset
 | 
236  | 
fun is_error (Token (_, (Error _, _), _)) = true  | 
| 
 
c197b3c3e7fa
some attempts to keep malformed syntax errors focussed, without too much red spilled onto the document view;
 
wenzelm 
parents: 
48743 
diff
changeset
 | 
237  | 
| is_error _ = false;  | 
| 
 
c197b3c3e7fa
some attempts to keep malformed syntax errors focussed, without too much red spilled onto the document view;
 
wenzelm 
parents: 
48743 
diff
changeset
 | 
238  | 
|
| 8651 | 239  | 
|
| 17069 | 240  | 
(* blanks and newlines -- space tokens obey lines *)  | 
| 8651 | 241  | 
|
| 
48771
 
2ea997196d04
clarified Command.range vs. Command.proper_range according to Scala version, which is potentially relevant for command status markup;
 
wenzelm 
parents: 
48764 
diff
changeset
 | 
242  | 
fun is_space (Token (_, (Space, _), _)) = true  | 
| 
 
2ea997196d04
clarified Command.range vs. Command.proper_range according to Scala version, which is potentially relevant for command status markup;
 
wenzelm 
parents: 
48764 
diff
changeset
 | 
243  | 
| is_space _ = false;  | 
| 
 
2ea997196d04
clarified Command.range vs. Command.proper_range according to Scala version, which is potentially relevant for command status markup;
 
wenzelm 
parents: 
48764 
diff
changeset
 | 
244  | 
|
| 27814 | 245  | 
fun is_blank (Token (_, (Space, x), _)) = not (String.isSuffix "\n" x)  | 
| 17069 | 246  | 
| is_blank _ = false;  | 
247  | 
||
| 27814 | 248  | 
fun is_newline (Token (_, (Space, x), _)) = String.isSuffix "\n" x  | 
| 8651 | 249  | 
| is_newline _ = false;  | 
250  | 
||
| 5825 | 251  | 
|
| 14991 | 252  | 
(* token content *)  | 
| 9155 | 253  | 
|
| 59795 | 254  | 
fun content_of (Token (_, (_, x), _)) = x;  | 
| 
25642
 
ebdff0dca2a5
text_of: made even more robust against recurrent errors;
 
wenzelm 
parents: 
25582 
diff
changeset
 | 
255  | 
|
| 59809 | 256  | 
fun input_of (Token ((source, range), (kind, _), _)) =  | 
| 59064 | 257  | 
Input.source (delimited_kind kind) source range;  | 
| 27873 | 258  | 
|
| 59795 | 259  | 
fun inner_syntax_of tok =  | 
260  | 
let val x = content_of tok  | 
|
| 59809 | 261  | 
in if YXML.detect x then x else Syntax.implode_input (input_of tok) end;  | 
| 
27747
 
d41abb7bc08a
token: maintain of source, which retains original position information;
 
wenzelm 
parents: 
27733 
diff
changeset
 | 
262  | 
|
| 
 
d41abb7bc08a
token: maintain of source, which retains original position information;
 
wenzelm 
parents: 
27733 
diff
changeset
 | 
263  | 
|
| 
55915
 
607948c90bf0
suppress short abbreviations more uniformly, for outer and quasi-outer syntax;
 
wenzelm 
parents: 
55914 
diff
changeset
 | 
264  | 
(* markup reports *)  | 
| 
55744
 
4a4e5686e091
clarified token markup: keyword1/keyword2 is for syntax, and "command" the entity kind;
 
wenzelm 
parents: 
55709 
diff
changeset
 | 
265  | 
|
| 
 
4a4e5686e091
clarified token markup: keyword1/keyword2 is for syntax, and "command" the entity kind;
 
wenzelm 
parents: 
55709 
diff
changeset
 | 
266  | 
local  | 
| 
 
4a4e5686e091
clarified token markup: keyword1/keyword2 is for syntax, and "command" the entity kind;
 
wenzelm 
parents: 
55709 
diff
changeset
 | 
267  | 
|
| 
 
4a4e5686e091
clarified token markup: keyword1/keyword2 is for syntax, and "command" the entity kind;
 
wenzelm 
parents: 
55709 
diff
changeset
 | 
268  | 
val token_kind_markup =  | 
| 59124 | 269  | 
fn Var => (Markup.var, "")  | 
| 59081 | 270  | 
| Type_Ident => (Markup.tfree, "")  | 
271  | 
| Type_Var => (Markup.tvar, "")  | 
|
272  | 
| String => (Markup.string, "")  | 
|
273  | 
| Alt_String => (Markup.alt_string, "")  | 
|
274  | 
| Verbatim => (Markup.verbatim, "")  | 
|
275  | 
| Cartouche => (Markup.cartouche, "")  | 
|
276  | 
| Comment => (Markup.comment, "")  | 
|
277  | 
| Error msg => (Markup.bad, msg)  | 
|
| 59124 | 278  | 
| _ => (Markup.empty, "");  | 
| 
55744
 
4a4e5686e091
clarified token markup: keyword1/keyword2 is for syntax, and "command" the entity kind;
 
wenzelm 
parents: 
55709 
diff
changeset
 | 
279  | 
|
| 59125 | 280  | 
fun keyword_reports tok = map (fn markup => ((pos_of tok, markup), ""));  | 
| 
59123
 
e68e44836d04
imitate command markup and rendering of Isabelle/jEdit in HTML output;
 
wenzelm 
parents: 
59112 
diff
changeset
 | 
281  | 
|
| 59125 | 282  | 
fun command_markups keywords x =  | 
283  | 
if Keyword.is_theory_end keywords x then [Markup.keyword2]  | 
|
284  | 
else if Keyword.is_proof_asm keywords x then [Markup.keyword3]  | 
|
285  | 
else if Keyword.is_improper keywords x then [Markup.keyword1, Markup.improper]  | 
|
286  | 
else [Markup.keyword1];  | 
|
| 
59123
 
e68e44836d04
imitate command markup and rendering of Isabelle/jEdit in HTML output;
 
wenzelm 
parents: 
59112 
diff
changeset
 | 
287  | 
|
| 56063 | 288  | 
in  | 
289  | 
||
| 56202 | 290  | 
fun keyword_markup (important, keyword) x =  | 
291  | 
if important orelse Symbol.is_ascii_identifier x then keyword else Markup.delimiter;  | 
|
| 
55919
 
2eb8c13339a5
more explicit quasi_keyword markup, for Args.$$$ material, which is somewhere in between of outer and inner syntax;
 
wenzelm 
parents: 
55916 
diff
changeset
 | 
292  | 
|
| 
55915
 
607948c90bf0
suppress short abbreviations more uniformly, for outer and quasi-outer syntax;
 
wenzelm 
parents: 
55914 
diff
changeset
 | 
293  | 
fun completion_report tok =  | 
| 
55914
 
c5b752d549e3
clarified init_assignable: make double-sure that initial values are reset;
 
wenzelm 
parents: 
55828 
diff
changeset
 | 
294  | 
if is_kind Keyword tok  | 
| 
55915
 
607948c90bf0
suppress short abbreviations more uniformly, for outer and quasi-outer syntax;
 
wenzelm 
parents: 
55914 
diff
changeset
 | 
295  | 
then map (fn m => ((pos_of tok, m), "")) (Completion.suppress_abbrevs (content_of tok))  | 
| 
 
607948c90bf0
suppress short abbreviations more uniformly, for outer and quasi-outer syntax;
 
wenzelm 
parents: 
55914 
diff
changeset
 | 
296  | 
else [];  | 
| 
55744
 
4a4e5686e091
clarified token markup: keyword1/keyword2 is for syntax, and "command" the entity kind;
 
wenzelm 
parents: 
55709 
diff
changeset
 | 
297  | 
|
| 59125 | 298  | 
fun reports keywords tok =  | 
| 
59123
 
e68e44836d04
imitate command markup and rendering of Isabelle/jEdit in HTML output;
 
wenzelm 
parents: 
59112 
diff
changeset
 | 
299  | 
if is_command tok then  | 
| 59125 | 300  | 
keyword_reports tok (command_markups keywords (content_of tok))  | 
| 
59123
 
e68e44836d04
imitate command markup and rendering of Isabelle/jEdit in HTML output;
 
wenzelm 
parents: 
59112 
diff
changeset
 | 
301  | 
else if is_kind Keyword tok then  | 
| 59125 | 302  | 
keyword_reports tok [keyword_markup (false, Markup.keyword2) (content_of tok)]  | 
| 
55915
 
607948c90bf0
suppress short abbreviations more uniformly, for outer and quasi-outer syntax;
 
wenzelm 
parents: 
55914 
diff
changeset
 | 
303  | 
else  | 
| 
 
607948c90bf0
suppress short abbreviations more uniformly, for outer and quasi-outer syntax;
 
wenzelm 
parents: 
55914 
diff
changeset
 | 
304  | 
let val (m, text) = token_kind_markup (kind_of tok)  | 
| 59125 | 305  | 
in [((pos_of tok, m), text)] end;  | 
| 
55915
 
607948c90bf0
suppress short abbreviations more uniformly, for outer and quasi-outer syntax;
 
wenzelm 
parents: 
55914 
diff
changeset
 | 
306  | 
|
| 59125 | 307  | 
fun markups keywords = map (#2 o #1) o reports keywords;  | 
| 
55914
 
c5b752d549e3
clarified init_assignable: make double-sure that initial values are reset;
 
wenzelm 
parents: 
55828 
diff
changeset
 | 
308  | 
|
| 
55744
 
4a4e5686e091
clarified token markup: keyword1/keyword2 is for syntax, and "command" the entity kind;
 
wenzelm 
parents: 
55709 
diff
changeset
 | 
309  | 
end;  | 
| 
 
4a4e5686e091
clarified token markup: keyword1/keyword2 is for syntax, and "command" the entity kind;
 
wenzelm 
parents: 
55709 
diff
changeset
 | 
310  | 
|
| 
 
4a4e5686e091
clarified token markup: keyword1/keyword2 is for syntax, and "command" the entity kind;
 
wenzelm 
parents: 
55709 
diff
changeset
 | 
311  | 
|
| 
27747
 
d41abb7bc08a
token: maintain of source, which retains original position information;
 
wenzelm 
parents: 
27733 
diff
changeset
 | 
312  | 
(* unparse *)  | 
| 
 
d41abb7bc08a
token: maintain of source, which retains original position information;
 
wenzelm 
parents: 
27733 
diff
changeset
 | 
313  | 
|
| 27814 | 314  | 
fun unparse (Token (_, (kind, x), _)) =  | 
| 14991 | 315  | 
(case kind of  | 
| 43773 | 316  | 
String => Symbol_Pos.quote_string_qq x  | 
| 59081 | 317  | 
| Alt_String => Symbol_Pos.quote_string_bq x  | 
| 43773 | 318  | 
  | Verbatim => enclose "{*" "*}" x
 | 
| 55033 | 319  | 
| Cartouche => cartouche x  | 
| 43773 | 320  | 
| Comment => enclose "(*" "*)" x  | 
| 
23729
 
d1ba656978c5
separated Malformed (symbolic char) from Error (bad input);
 
wenzelm 
parents: 
23721 
diff
changeset
 | 
321  | 
| EOF => ""  | 
| 14991 | 322  | 
| _ => x);  | 
323  | 
||
| 59125 | 324  | 
fun print tok = Markup.markups (markups Keyword.empty_keywords tok) (unparse tok);  | 
| 55745 | 325  | 
|
| 
23788
 
54ce229dc858
Symbol.not_eof/sync is superceded by Symbol.is_regular (rules out further control symbols);
 
wenzelm 
parents: 
23729 
diff
changeset
 | 
326  | 
fun text_of tok =  | 
| 
58861
 
5ff61774df11
command-line terminator ";" is no longer accepted;
 
wenzelm 
parents: 
58855 
diff
changeset
 | 
327  | 
let  | 
| 
 
5ff61774df11
command-line terminator ";" is no longer accepted;
 
wenzelm 
parents: 
58855 
diff
changeset
 | 
328  | 
val k = str_of_kind (kind_of tok);  | 
| 59125 | 329  | 
val ms = markups Keyword.empty_keywords tok;  | 
| 
58861
 
5ff61774df11
command-line terminator ";" is no longer accepted;
 
wenzelm 
parents: 
58855 
diff
changeset
 | 
330  | 
val s = unparse tok;  | 
| 
 
5ff61774df11
command-line terminator ";" is no longer accepted;
 
wenzelm 
parents: 
58855 
diff
changeset
 | 
331  | 
in  | 
| 
 
5ff61774df11
command-line terminator ";" is no longer accepted;
 
wenzelm 
parents: 
58855 
diff
changeset
 | 
332  | 
if s = "" then (k, "")  | 
| 
 
5ff61774df11
command-line terminator ";" is no longer accepted;
 
wenzelm 
parents: 
58855 
diff
changeset
 | 
333  | 
else if size s < 40 andalso not (exists_string (fn c => c = "\n") s)  | 
| 59125 | 334  | 
then (k ^ " " ^ Markup.markups ms s, "")  | 
335  | 
else (k, Markup.markups ms s)  | 
|
| 
58861
 
5ff61774df11
command-line terminator ";" is no longer accepted;
 
wenzelm 
parents: 
58855 
diff
changeset
 | 
336  | 
end;  | 
| 
23729
 
d1ba656978c5
separated Malformed (symbolic char) from Error (bad input);
 
wenzelm 
parents: 
23721 
diff
changeset
 | 
337  | 
|
| 5825 | 338  | 
|
339  | 
||
| 27814 | 340  | 
(** associated values **)  | 
341  | 
||
| 
48867
 
e9beabf045ab
some support for inlining file content into outer syntax token language;
 
wenzelm 
parents: 
48771 
diff
changeset
 | 
342  | 
(* inlined file content *)  | 
| 
 
e9beabf045ab
some support for inlining file content into outer syntax token language;
 
wenzelm 
parents: 
48771 
diff
changeset
 | 
343  | 
|
| 
54519
 
5fed81762406
maintain blobs within document state: digest + text in ML, digest-only in Scala;
 
wenzelm 
parents: 
51266 
diff
changeset
 | 
344  | 
fun get_files (Token (_, _, Value (SOME (Files files)))) = files  | 
| 
 
5fed81762406
maintain blobs within document state: digest + text in ML, digest-only in Scala;
 
wenzelm 
parents: 
51266 
diff
changeset
 | 
345  | 
| get_files _ = [];  | 
| 
48867
 
e9beabf045ab
some support for inlining file content into outer syntax token language;
 
wenzelm 
parents: 
48771 
diff
changeset
 | 
346  | 
|
| 
54519
 
5fed81762406
maintain blobs within document state: digest + text in ML, digest-only in Scala;
 
wenzelm 
parents: 
51266 
diff
changeset
 | 
347  | 
fun put_files [] tok = tok  | 
| 
 
5fed81762406
maintain blobs within document state: digest + text in ML, digest-only in Scala;
 
wenzelm 
parents: 
51266 
diff
changeset
 | 
348  | 
| put_files files (Token (x, y, Slot)) = Token (x, y, Value (SOME (Files files)))  | 
| 55708 | 349  | 
  | put_files _ tok = raise Fail ("Cannot put inlined files here" ^ Position.here (pos_of tok));
 | 
| 
48867
 
e9beabf045ab
some support for inlining file content into outer syntax token language;
 
wenzelm 
parents: 
48771 
diff
changeset
 | 
350  | 
|
| 
 
e9beabf045ab
some support for inlining file content into outer syntax token language;
 
wenzelm 
parents: 
48771 
diff
changeset
 | 
351  | 
|
| 27814 | 352  | 
(* access values *)  | 
353  | 
||
354  | 
fun get_value (Token (_, _, Value v)) = v  | 
|
355  | 
| get_value _ = NONE;  | 
|
356  | 
||
357  | 
fun map_value f (Token (x, y, Value (SOME v))) = Token (x, y, Value (SOME (f v)))  | 
|
358  | 
| map_value _ tok = tok;  | 
|
359  | 
||
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
360  | 
|
| 61822 | 361  | 
(* reports of value *)  | 
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
362  | 
|
| 61822 | 363  | 
fun get_assignable_value (Token (_, _, Assignable r)) = ! r  | 
364  | 
| get_assignable_value (Token (_, _, Value v)) = v  | 
|
365  | 
| get_assignable_value _ = NONE;  | 
|
| 60211 | 366  | 
|
| 
55914
 
c5b752d549e3
clarified init_assignable: make double-sure that initial values are reset;
 
wenzelm 
parents: 
55828 
diff
changeset
 | 
367  | 
fun reports_of_value tok =  | 
| 60211 | 368  | 
(case get_assignable_value tok of  | 
| 56063 | 369  | 
SOME (Literal markup) =>  | 
370  | 
let  | 
|
371  | 
val pos = pos_of tok;  | 
|
372  | 
val x = content_of tok;  | 
|
373  | 
in  | 
|
374  | 
if Position.is_reported pos then  | 
|
375  | 
map (pair pos) (keyword_markup markup x :: Completion.suppress_abbrevs x)  | 
|
376  | 
else []  | 
|
377  | 
end  | 
|
378  | 
| _ => []);  | 
|
| 
55914
 
c5b752d549e3
clarified init_assignable: make double-sure that initial values are reset;
 
wenzelm 
parents: 
55828 
diff
changeset
 | 
379  | 
|
| 27814 | 380  | 
|
| 61822 | 381  | 
(* name value *)  | 
382  | 
||
383  | 
fun name_value a = Name (a, Morphism.identity);  | 
|
384  | 
||
385  | 
fun get_name tok =  | 
|
386  | 
(case get_assignable_value tok of  | 
|
387  | 
SOME (Name (a, _)) => SOME a  | 
|
388  | 
| _ => NONE);  | 
|
389  | 
||
390  | 
||
| 59646 | 391  | 
(* maxidx *)  | 
392  | 
||
393  | 
fun declare_maxidx tok =  | 
|
394  | 
(case get_value tok of  | 
|
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
395  | 
SOME (Source src) => fold declare_maxidx src  | 
| 59646 | 396  | 
| SOME (Typ T) => Variable.declare_maxidx (Term.maxidx_of_typ T)  | 
397  | 
| SOME (Term t) => Variable.declare_maxidx (Term.maxidx_of_term t)  | 
|
398  | 
| SOME (Fact (_, ths)) => fold (Variable.declare_maxidx o Thm.maxidx_of) ths  | 
|
399  | 
| SOME (Attribute _) => I (* FIXME !? *)  | 
|
400  | 
| SOME (Declaration decl) =>  | 
|
401  | 
(fn ctxt =>  | 
|
402  | 
let val ctxt' = Context.proof_map (Morphism.form decl) ctxt  | 
|
403  | 
in Variable.declare_maxidx (Variable.maxidx_of ctxt') ctxt end)  | 
|
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
404  | 
| _ => I);  | 
| 59646 | 405  | 
|
406  | 
||
| 61080 | 407  | 
(* fact values *)  | 
408  | 
||
409  | 
fun map_facts f =  | 
|
410  | 
map_value (fn v =>  | 
|
411  | 
(case v of  | 
|
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
412  | 
Source src => Source (map (map_facts f) src)  | 
| 61820 | 413  | 
| Fact (a, ths) => Fact (a, f a ths)  | 
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
414  | 
| _ => v));  | 
| 61080 | 415  | 
|
416  | 
||
| 58012 | 417  | 
(* transform *)  | 
418  | 
||
419  | 
fun transform phi =  | 
|
420  | 
map_value (fn v =>  | 
|
421  | 
(case v of  | 
|
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
422  | 
Source src => Source (map (transform phi) src)  | 
| 58012 | 423  | 
| Literal _ => v  | 
424  | 
| Name (a, psi) => Name (a, psi $> phi)  | 
|
425  | 
| Typ T => Typ (Morphism.typ phi T)  | 
|
426  | 
| Term t => Term (Morphism.term phi t)  | 
|
427  | 
| Fact (a, ths) => Fact (a, Morphism.fact phi ths)  | 
|
428  | 
| Attribute att => Attribute (Morphism.transform phi att)  | 
|
| 58017 | 429  | 
| Declaration decl => Declaration (Morphism.transform phi decl)  | 
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
430  | 
| Files _ => v));  | 
| 58012 | 431  | 
|
432  | 
||
433  | 
(* static binding *)  | 
|
434  | 
||
435  | 
(*1st stage: initialize assignable slots*)  | 
|
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
436  | 
fun init_assignable tok =  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
437  | 
(case tok of  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
438  | 
Token (x, y, Slot) => Token (x, y, Assignable (Unsynchronized.ref NONE))  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
439  | 
| Token (_, _, Value _) => tok  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
440  | 
| Token (_, _, Assignable r) => (r := NONE; tok));  | 
| 58012 | 441  | 
|
442  | 
(*2nd stage: assign values as side-effect of scanning*)  | 
|
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
443  | 
fun assign v tok =  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
444  | 
(case tok of  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
445  | 
Token (x, y, Slot) => Token (x, y, Value v)  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
446  | 
| Token (_, _, Value _) => tok  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
447  | 
| Token (_, _, Assignable r) => (r := v; tok));  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
448  | 
|
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
449  | 
fun evaluate mk eval arg =  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
450  | 
let val x = eval arg in (assign (SOME (mk x)) arg; x) end;  | 
| 58012 | 451  | 
|
452  | 
(*3rd stage: static closure of final values*)  | 
|
453  | 
fun closure (Token (x, y, Assignable (Unsynchronized.ref v))) = Token (x, y, Value v)  | 
|
454  | 
| closure tok = tok;  | 
|
455  | 
||
456  | 
||
457  | 
(* pretty *)  | 
|
| 
57944
 
fff8d328da56
more informative Token.Name with history of morphisms;
 
wenzelm 
parents: 
57942 
diff
changeset
 | 
458  | 
|
| 
 
fff8d328da56
more informative Token.Name with history of morphisms;
 
wenzelm 
parents: 
57942 
diff
changeset
 | 
459  | 
fun pretty_value ctxt tok =  | 
| 
 
fff8d328da56
more informative Token.Name with history of morphisms;
 
wenzelm 
parents: 
57942 
diff
changeset
 | 
460  | 
(case get_value tok of  | 
| 
 
fff8d328da56
more informative Token.Name with history of morphisms;
 
wenzelm 
parents: 
57942 
diff
changeset
 | 
461  | 
SOME (Literal markup) =>  | 
| 
 
fff8d328da56
more informative Token.Name with history of morphisms;
 
wenzelm 
parents: 
57942 
diff
changeset
 | 
462  | 
let val x = content_of tok  | 
| 
 
fff8d328da56
more informative Token.Name with history of morphisms;
 
wenzelm 
parents: 
57942 
diff
changeset
 | 
463  | 
in Pretty.mark_str (keyword_markup markup x, x) end  | 
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
464  | 
  | SOME (Name ({print, ...}, _)) => Pretty.quote (Pretty.mark_str (print ctxt))
 | 
| 
57944
 
fff8d328da56
more informative Token.Name with history of morphisms;
 
wenzelm 
parents: 
57942 
diff
changeset
 | 
465  | 
| SOME (Typ T) => Syntax.pretty_typ ctxt T  | 
| 
 
fff8d328da56
more informative Token.Name with history of morphisms;
 
wenzelm 
parents: 
57942 
diff
changeset
 | 
466  | 
| SOME (Term t) => Syntax.pretty_term ctxt t  | 
| 
 
fff8d328da56
more informative Token.Name with history of morphisms;
 
wenzelm 
parents: 
57942 
diff
changeset
 | 
467  | 
| SOME (Fact (_, ths)) =>  | 
| 62094 | 468  | 
      Pretty.enclose "(" ")" (Pretty.breaks (map (Pretty.cartouche o Thm.pretty_thm ctxt) ths))
 | 
| 59125 | 469  | 
| _ => Pretty.marks_str (markups Keyword.empty_keywords tok, unparse tok));  | 
| 
57944
 
fff8d328da56
more informative Token.Name with history of morphisms;
 
wenzelm 
parents: 
57942 
diff
changeset
 | 
470  | 
|
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
471  | 
|
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
472  | 
(* src *)  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
473  | 
|
| 61825 | 474  | 
fun dest_src ([]: src) = raise Fail "Empty token source"  | 
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
475  | 
| dest_src (head :: args) = (head, args);  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
476  | 
|
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
477  | 
fun name_of_src src =  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
478  | 
let  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
479  | 
val head = #1 (dest_src src);  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
480  | 
val name =  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
481  | 
(case get_name head of  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
482  | 
        SOME {name, ...} => name
 | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
483  | 
| NONE => content_of head);  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
484  | 
in (name, pos_of head) end;  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
485  | 
|
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
486  | 
val args_of_src = #2 o dest_src;  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
487  | 
|
| 
58011
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57944 
diff
changeset
 | 
488  | 
fun pretty_src ctxt src =  | 
| 
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57944 
diff
changeset
 | 
489  | 
let  | 
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
490  | 
val (head, args) = dest_src src;  | 
| 
58011
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57944 
diff
changeset
 | 
491  | 
val prt_name =  | 
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
492  | 
(case get_name head of  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
493  | 
        SOME {print, ...} => Pretty.mark_str (print ctxt)
 | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
494  | 
| NONE => Pretty.str (content_of head));  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
495  | 
in Pretty.block (Pretty.breaks (Pretty.quote prt_name :: map (pretty_value ctxt) args)) end;  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
496  | 
|
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
497  | 
fun checked_src (head :: _) = is_some (get_name head)  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
498  | 
| checked_src [] = true;  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
499  | 
|
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
500  | 
fun check_src ctxt get_table src =  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
501  | 
let  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
502  | 
val (head, args) = dest_src src;  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
503  | 
val table = get_table ctxt;  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
504  | 
in  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
505  | 
(case get_name head of  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
506  | 
      SOME {name, ...} => (src, Name_Space.get table name)
 | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
507  | 
| NONE =>  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
508  | 
let  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
509  | 
val (name, x) =  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
510  | 
Name_Space.check (Context.Proof ctxt) table (content_of head, pos_of head);  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
511  | 
val kind = Name_Space.kind_of (Name_Space.space_of_table table);  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
512  | 
fun print ctxt' =  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
513  | 
Name_Space.markup_extern ctxt' (Name_Space.space_of_table (get_table ctxt')) name;  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
514  | 
          val value = name_value {name = name, kind = kind, print = print};
 | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
515  | 
val head' = closure (assign (SOME value) head);  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
516  | 
in (head' :: args, x) end)  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
517  | 
end;  | 
| 
58011
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57944 
diff
changeset
 | 
518  | 
|
| 
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57944 
diff
changeset
 | 
519  | 
|
| 59125 | 520  | 
|
| 5825 | 521  | 
(** scanners **)  | 
522  | 
||
| 30573 | 523  | 
open Basic_Symbol_Pos;  | 
| 5825 | 524  | 
|
| 48764 | 525  | 
val err_prefix = "Outer lexical error: ";  | 
526  | 
||
527  | 
fun !!! msg = Symbol_Pos.!!! (fn () => err_prefix ^ msg);  | 
|
| 5825 | 528  | 
|
529  | 
||
530  | 
(* scan symbolic idents *)  | 
|
531  | 
||
| 8231 | 532  | 
val scan_symid =  | 
| 55033 | 533  | 
Scan.many1 (Symbol.is_symbolic_char o Symbol_Pos.symbol) ||  | 
| 
40525
 
14a2e686bdac
eliminated slightly odd pervasive Symbol_Pos.symbol;
 
wenzelm 
parents: 
40523 
diff
changeset
 | 
534  | 
Scan.one (Symbol.is_symbolic o Symbol_Pos.symbol) >> single;  | 
| 5825 | 535  | 
|
| 8231 | 536  | 
fun is_symid str =  | 
537  | 
(case try Symbol.explode str of  | 
|
| 55033 | 538  | 
SOME [s] => Symbol.is_symbolic s orelse Symbol.is_symbolic_char s  | 
539  | 
| SOME ss => forall Symbol.is_symbolic_char ss  | 
|
| 8231 | 540  | 
| _ => false);  | 
541  | 
||
| 27814 | 542  | 
fun ident_or_symbolic "begin" = false  | 
543  | 
| ident_or_symbolic ":" = true  | 
|
544  | 
| ident_or_symbolic "::" = true  | 
|
| 50239 | 545  | 
| ident_or_symbolic s = Symbol_Pos.is_identifier s orelse is_symid s;  | 
| 5825 | 546  | 
|
547  | 
||
548  | 
(* scan verbatim text *)  | 
|
549  | 
||
550  | 
val scan_verb =  | 
|
| 55107 | 551  | 
$$$ "*" --| Scan.ahead (~$$ "}") ||  | 
| 58854 | 552  | 
Scan.one (fn (s, _) => s <> "*" andalso Symbol.not_eof s) >> single;  | 
| 5825 | 553  | 
|
554  | 
val scan_verbatim =  | 
|
| 55107 | 555  | 
  Scan.ahead ($$ "{" -- $$ "*") |--
 | 
| 55106 | 556  | 
!!! "unclosed verbatim text"  | 
| 55107 | 557  | 
      ((Symbol_Pos.scan_pos --| $$ "{" --| $$ "*") --
 | 
| 61476 | 558  | 
(Scan.repeats scan_verb -- ($$ "*" |-- $$ "}" |-- Symbol_Pos.scan_pos)));  | 
| 5825 | 559  | 
|
| 
48743
 
a72f8ffecf31
refined recovery of scan errors: longest prefix of delimited token after failure, otherwise just one symbol;
 
wenzelm 
parents: 
48741 
diff
changeset
 | 
560  | 
val recover_verbatim =  | 
| 61476 | 561  | 
  $$$ "{" @@@ $$$ "*" @@@ Scan.repeats scan_verb;
 | 
| 
48743
 
a72f8ffecf31
refined recovery of scan errors: longest prefix of delimited token after failure, otherwise just one symbol;
 
wenzelm 
parents: 
48741 
diff
changeset
 | 
562  | 
|
| 5825 | 563  | 
|
| 55033 | 564  | 
(* scan cartouche *)  | 
565  | 
||
566  | 
val scan_cartouche =  | 
|
| 
55104
 
8284c0d5bf52
clarified scan_cartouche_depth, according to Scala version;
 
wenzelm 
parents: 
55103 
diff
changeset
 | 
567  | 
Symbol_Pos.scan_pos --  | 
| 55105 | 568  | 
((Symbol_Pos.scan_cartouche err_prefix >> Symbol_Pos.cartouche_content) -- Symbol_Pos.scan_pos);  | 
| 55033 | 569  | 
|
570  | 
||
| 5825 | 571  | 
(* scan space *)  | 
572  | 
||
| 
48771
 
2ea997196d04
clarified Command.range vs. Command.proper_range according to Scala version, which is potentially relevant for command status markup;
 
wenzelm 
parents: 
48764 
diff
changeset
 | 
573  | 
fun space_symbol (s, _) = Symbol.is_blank s andalso s <> "\n";  | 
| 5825 | 574  | 
|
575  | 
val scan_space =  | 
|
| 
48771
 
2ea997196d04
clarified Command.range vs. Command.proper_range according to Scala version, which is potentially relevant for command status markup;
 
wenzelm 
parents: 
48764 
diff
changeset
 | 
576  | 
Scan.many1 space_symbol @@@ Scan.optional ($$$ "\n") [] ||  | 
| 
 
2ea997196d04
clarified Command.range vs. Command.proper_range according to Scala version, which is potentially relevant for command status markup;
 
wenzelm 
parents: 
48764 
diff
changeset
 | 
577  | 
Scan.many space_symbol @@@ $$$ "\n";  | 
| 5825 | 578  | 
|
579  | 
||
| 
27780
 
7d0910f662f7
more precise positions due to SymbolsPos.implode_delim;
 
wenzelm 
parents: 
27769 
diff
changeset
 | 
580  | 
(* scan comment *)  | 
| 5825 | 581  | 
|
582  | 
val scan_comment =  | 
|
| 55105 | 583  | 
Symbol_Pos.scan_pos -- (Symbol_Pos.scan_comment_body err_prefix -- Symbol_Pos.scan_pos);  | 
| 5825 | 584  | 
|
585  | 
||
| 27663 | 586  | 
|
| 27769 | 587  | 
(** token sources **)  | 
| 5825 | 588  | 
|
| 27769 | 589  | 
fun source_proper src = src |> Source.filter is_proper;  | 
| 5825 | 590  | 
|
| 23678 | 591  | 
local  | 
592  | 
||
| 27769 | 593  | 
fun token_leq ((_, syms1), (_, syms2)) = length syms1 <= length syms2;  | 
| 
27780
 
7d0910f662f7
more precise positions due to SymbolsPos.implode_delim;
 
wenzelm 
parents: 
27769 
diff
changeset
 | 
594  | 
|
| 27799 | 595  | 
fun token k ss =  | 
| 
43709
 
717e96cf9527
discontinued special treatment of hard tabulators;
 
wenzelm 
parents: 
42503 
diff
changeset
 | 
596  | 
Token ((Symbol_Pos.implode ss, Symbol_Pos.range ss), (k, Symbol_Pos.content ss), Slot);  | 
| 27799 | 597  | 
|
598  | 
fun token_range k (pos1, (ss, pos2)) =  | 
|
| 59112 | 599  | 
Token (Symbol_Pos.implode_range (pos1, pos2) ss, (k, Symbol_Pos.content ss), Slot);  | 
| 23678 | 600  | 
|
| 59083 | 601  | 
fun scan_token keywords = !!! "bad input"  | 
| 48764 | 602  | 
(Symbol_Pos.scan_string_qq err_prefix >> token_range String ||  | 
| 59081 | 603  | 
Symbol_Pos.scan_string_bq err_prefix >> token_range Alt_String ||  | 
| 27799 | 604  | 
scan_verbatim >> token_range Verbatim ||  | 
| 55033 | 605  | 
scan_cartouche >> token_range Cartouche ||  | 
| 27799 | 606  | 
scan_comment >> token_range Comment ||  | 
| 
27780
 
7d0910f662f7
more precise positions due to SymbolsPos.implode_delim;
 
wenzelm 
parents: 
27769 
diff
changeset
 | 
607  | 
scan_space >> token Space ||  | 
| 
 
7d0910f662f7
more precise positions due to SymbolsPos.implode_delim;
 
wenzelm 
parents: 
27769 
diff
changeset
 | 
608  | 
(Scan.max token_leq  | 
| 27769 | 609  | 
(Scan.max token_leq  | 
| 58903 | 610  | 
(Scan.literal (Keyword.major_keywords keywords) >> pair Command)  | 
611  | 
(Scan.literal (Keyword.minor_keywords keywords) >> pair Keyword))  | 
|
| 59081 | 612  | 
(Lexicon.scan_longid >> pair Long_Ident ||  | 
| 
42290
 
b1f544c84040
discontinued special treatment of structure Lexicon;
 
wenzelm 
parents: 
40958 
diff
changeset
 | 
613  | 
Lexicon.scan_id >> pair Ident ||  | 
| 
 
b1f544c84040
discontinued special treatment of structure Lexicon;
 
wenzelm 
parents: 
40958 
diff
changeset
 | 
614  | 
Lexicon.scan_var >> pair Var ||  | 
| 59081 | 615  | 
Lexicon.scan_tid >> pair Type_Ident ||  | 
616  | 
Lexicon.scan_tvar >> pair Type_Var ||  | 
|
| 62782 | 617  | 
Symbol_Pos.scan_float >> pair Float ||  | 
618  | 
Symbol_Pos.scan_nat >> pair Nat ||  | 
|
| 59081 | 619  | 
scan_symid >> pair Sym_Ident) >> uncurry token));  | 
| 27769 | 620  | 
|
621  | 
fun recover msg =  | 
|
| 
48743
 
a72f8ffecf31
refined recovery of scan errors: longest prefix of delimited token after failure, otherwise just one symbol;
 
wenzelm 
parents: 
48741 
diff
changeset
 | 
622  | 
(Symbol_Pos.recover_string_qq ||  | 
| 
 
a72f8ffecf31
refined recovery of scan errors: longest prefix of delimited token after failure, otherwise just one symbol;
 
wenzelm 
parents: 
48741 
diff
changeset
 | 
623  | 
Symbol_Pos.recover_string_bq ||  | 
| 
 
a72f8ffecf31
refined recovery of scan errors: longest prefix of delimited token after failure, otherwise just one symbol;
 
wenzelm 
parents: 
48741 
diff
changeset
 | 
624  | 
recover_verbatim ||  | 
| 55033 | 625  | 
Symbol_Pos.recover_cartouche ||  | 
| 
48743
 
a72f8ffecf31
refined recovery of scan errors: longest prefix of delimited token after failure, otherwise just one symbol;
 
wenzelm 
parents: 
48741 
diff
changeset
 | 
626  | 
Symbol_Pos.recover_comment ||  | 
| 58854 | 627  | 
Scan.one (Symbol.not_eof o Symbol_Pos.symbol) >> single)  | 
| 
27780
 
7d0910f662f7
more precise positions due to SymbolsPos.implode_delim;
 
wenzelm 
parents: 
27769 
diff
changeset
 | 
628  | 
>> (single o token (Error msg));  | 
| 23678 | 629  | 
|
630  | 
in  | 
|
| 5825 | 631  | 
|
| 
58904
 
f49c4f883c58
eliminated pointless dynamic keywords (TTY legacy);
 
wenzelm 
parents: 
58903 
diff
changeset
 | 
632  | 
fun source' strict keywords =  | 
| 58864 | 633  | 
let  | 
| 59083 | 634  | 
val scan_strict = Scan.bulk (scan_token keywords);  | 
| 58864 | 635  | 
val scan = if strict then scan_strict else Scan.recover scan_strict recover;  | 
636  | 
in Source.source Symbol_Pos.stopper scan end;  | 
|
| 
27780
 
7d0910f662f7
more precise positions due to SymbolsPos.implode_delim;
 
wenzelm 
parents: 
27769 
diff
changeset
 | 
637  | 
|
| 
58904
 
f49c4f883c58
eliminated pointless dynamic keywords (TTY legacy);
 
wenzelm 
parents: 
58903 
diff
changeset
 | 
638  | 
fun source keywords pos src = Symbol_Pos.source pos src |> source' false keywords;  | 
| 
 
f49c4f883c58
eliminated pointless dynamic keywords (TTY legacy);
 
wenzelm 
parents: 
58903 
diff
changeset
 | 
639  | 
fun source_strict keywords pos src = Symbol_Pos.source pos src |> source' true keywords;  | 
| 23678 | 640  | 
|
| 61471 | 641  | 
fun read_cartouche syms =  | 
642  | 
(case Scan.read Symbol_Pos.stopper (scan_cartouche >> token_range Cartouche) syms of  | 
|
643  | 
SOME tok => tok  | 
|
644  | 
  | NONE => error ("Single cartouche expected" ^ Position.here (#1 (Symbol_Pos.range syms))));
 | 
|
645  | 
||
| 23678 | 646  | 
end;  | 
| 5825 | 647  | 
|
| 
30586
 
9674f64a0702
moved basic change_prompt, scan_string, scan_alt_string, scan_quoted to symbol_pos.ML;
 
wenzelm 
parents: 
30573 
diff
changeset
 | 
648  | 
|
| 59083 | 649  | 
(* explode *)  | 
650  | 
||
651  | 
fun explode keywords pos =  | 
|
652  | 
Source.of_string #>  | 
|
653  | 
Symbol.source #>  | 
|
654  | 
source keywords pos #>  | 
|
655  | 
Source.exhaust;  | 
|
656  | 
||
657  | 
||
| 
59085
 
08a6901eb035
clarified define_command: send tokens more directly, without requiring keywords in ML;
 
wenzelm 
parents: 
59083 
diff
changeset
 | 
658  | 
(* make *)  | 
| 
 
08a6901eb035
clarified define_command: send tokens more directly, without requiring keywords in ML;
 
wenzelm 
parents: 
59083 
diff
changeset
 | 
659  | 
|
| 
 
08a6901eb035
clarified define_command: send tokens more directly, without requiring keywords in ML;
 
wenzelm 
parents: 
59083 
diff
changeset
 | 
660  | 
fun make ((k, n), s) pos =  | 
| 
 
08a6901eb035
clarified define_command: send tokens more directly, without requiring keywords in ML;
 
wenzelm 
parents: 
59083 
diff
changeset
 | 
661  | 
let  | 
| 
 
08a6901eb035
clarified define_command: send tokens more directly, without requiring keywords in ML;
 
wenzelm 
parents: 
59083 
diff
changeset
 | 
662  | 
val pos' = Position.advance_offset n pos;  | 
| 62797 | 663  | 
val range = Position.range (pos, pos');  | 
| 
59085
 
08a6901eb035
clarified define_command: send tokens more directly, without requiring keywords in ML;
 
wenzelm 
parents: 
59083 
diff
changeset
 | 
664  | 
val tok =  | 
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
665  | 
if 0 <= k andalso k < Vector.length immediate_kinds then  | 
| 
59085
 
08a6901eb035
clarified define_command: send tokens more directly, without requiring keywords in ML;
 
wenzelm 
parents: 
59083 
diff
changeset
 | 
666  | 
Token ((s, range), (Vector.sub (immediate_kinds, k), s), Slot)  | 
| 
 
08a6901eb035
clarified define_command: send tokens more directly, without requiring keywords in ML;
 
wenzelm 
parents: 
59083 
diff
changeset
 | 
667  | 
else  | 
| 
 
08a6901eb035
clarified define_command: send tokens more directly, without requiring keywords in ML;
 
wenzelm 
parents: 
59083 
diff
changeset
 | 
668  | 
(case explode Keyword.empty_keywords pos s of  | 
| 
 
08a6901eb035
clarified define_command: send tokens more directly, without requiring keywords in ML;
 
wenzelm 
parents: 
59083 
diff
changeset
 | 
669  | 
[tok] => tok  | 
| 
 
08a6901eb035
clarified define_command: send tokens more directly, without requiring keywords in ML;
 
wenzelm 
parents: 
59083 
diff
changeset
 | 
670  | 
| _ => Token ((s, range), (Error (err_prefix ^ "exactly one token expected"), s), Slot))  | 
| 
 
08a6901eb035
clarified define_command: send tokens more directly, without requiring keywords in ML;
 
wenzelm 
parents: 
59083 
diff
changeset
 | 
671  | 
in (tok, pos') end;  | 
| 
 
08a6901eb035
clarified define_command: send tokens more directly, without requiring keywords in ML;
 
wenzelm 
parents: 
59083 
diff
changeset
 | 
672  | 
|
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
673  | 
fun make_string (s, pos) =  | 
| 62799 | 674  | 
let  | 
675  | 
val Token ((x, _), y, z) = #1 (make ((~1, 0), Symbol_Pos.quote_string_qq s) Position.none);  | 
|
| 62800 | 676  | 
val pos' = Position.no_range_position pos;  | 
| 62799 | 677  | 
in Token ((x, (pos', pos')), y, z) end;  | 
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
678  | 
|
| 63019 | 679  | 
val make_int = explode Keyword.empty_keywords Position.none o signed_string_of_int;  | 
680  | 
||
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
681  | 
fun make_src a args = make_string a :: args;  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
682  | 
|
| 
59085
 
08a6901eb035
clarified define_command: send tokens more directly, without requiring keywords in ML;
 
wenzelm 
parents: 
59083 
diff
changeset
 | 
683  | 
|
| 
58011
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57944 
diff
changeset
 | 
684  | 
|
| 
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57944 
diff
changeset
 | 
685  | 
(** parsers **)  | 
| 
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57944 
diff
changeset
 | 
686  | 
|
| 
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57944 
diff
changeset
 | 
687  | 
type 'a parser = T list -> 'a * T list;  | 
| 
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57944 
diff
changeset
 | 
688  | 
type 'a context_parser = Context.generic * T list -> 'a * (Context.generic * T list);  | 
| 
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57944 
diff
changeset
 | 
689  | 
|
| 58045 | 690  | 
|
| 61471 | 691  | 
(* read antiquotation source *)  | 
| 58045 | 692  | 
|
| 58903 | 693  | 
fun read_no_commands keywords scan syms =  | 
| 58045 | 694  | 
Source.of_list syms  | 
| 
58904
 
f49c4f883c58
eliminated pointless dynamic keywords (TTY legacy);
 
wenzelm 
parents: 
58903 
diff
changeset
 | 
695  | 
|> source' true (Keyword.no_command_keywords keywords)  | 
| 58045 | 696  | 
|> source_proper  | 
| 58864 | 697  | 
|> Source.source stopper (Scan.error (Scan.bulk scan))  | 
| 58045 | 698  | 
|> Source.exhaust;  | 
699  | 
||
| 58903 | 700  | 
fun read_antiq keywords scan (syms, pos) =  | 
| 58045 | 701  | 
let  | 
702  | 
fun err msg =  | 
|
703  | 
      cat_error msg ("Malformed antiquotation" ^ Position.here pos ^ ":\n" ^
 | 
|
704  | 
        "@{" ^ Symbol_Pos.content syms ^ "}");
 | 
|
| 58903 | 705  | 
val res = read_no_commands keywords scan syms handle ERROR msg => err msg;  | 
| 58045 | 706  | 
in (case res of [x] => x | _ => err "") end;  | 
707  | 
||
708  | 
||
709  | 
(* wrapped syntax *)  | 
|
710  | 
||
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
711  | 
fun syntax_generic scan src context =  | 
| 
58011
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57944 
diff
changeset
 | 
712  | 
let  | 
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
713  | 
val (name, pos) = name_of_src src;  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
714  | 
val args1 = map init_assignable (args_of_src src);  | 
| 
58011
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57944 
diff
changeset
 | 
715  | 
fun reported_text () =  | 
| 
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57944 
diff
changeset
 | 
716  | 
if Context_Position.is_visible_generic context then  | 
| 
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57944 
diff
changeset
 | 
717  | 
((pos, Markup.operator) :: maps (reports_of_value o closure) args1)  | 
| 
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57944 
diff
changeset
 | 
718  | 
|> map (fn (p, m) => Position.reported_text p m "")  | 
| 
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57944 
diff
changeset
 | 
719  | 
else [];  | 
| 
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57944 
diff
changeset
 | 
720  | 
in  | 
| 
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57944 
diff
changeset
 | 
721  | 
(case Scan.error (Scan.finite' stopper (Scan.option scan)) (context, args1) of  | 
| 
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57944 
diff
changeset
 | 
722  | 
(SOME x, (context', [])) =>  | 
| 
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57944 
diff
changeset
 | 
723  | 
let val _ = Output.report (reported_text ())  | 
| 
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57944 
diff
changeset
 | 
724  | 
in (x, context') end  | 
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
725  | 
| (_, (context', args2)) =>  | 
| 
58011
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57944 
diff
changeset
 | 
726  | 
let  | 
| 
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57944 
diff
changeset
 | 
727  | 
val print_name =  | 
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
728  | 
(case get_name (hd src) of  | 
| 
58011
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57944 
diff
changeset
 | 
729  | 
NONE => quote name  | 
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
730  | 
            | SOME {kind, print, ...} =>
 | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
731  | 
let  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
732  | 
val ctxt' = Context.proof_of context';  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
733  | 
val (markup, xname) = print ctxt';  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61614 
diff
changeset
 | 
734  | 
in plain_words kind ^ " " ^ quote (Markup.markup markup xname) end);  | 
| 
58011
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57944 
diff
changeset
 | 
735  | 
val print_args =  | 
| 
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57944 
diff
changeset
 | 
736  | 
if null args2 then "" else ":\n " ^ space_implode " " (map print args2);  | 
| 
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57944 
diff
changeset
 | 
737  | 
in  | 
| 
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57944 
diff
changeset
 | 
738  | 
          error ("Bad arguments for " ^ print_name ^ Position.here pos ^ print_args ^
 | 
| 
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57944 
diff
changeset
 | 
739  | 
Markup.markup_report (implode (reported_text ())))  | 
| 
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57944 
diff
changeset
 | 
740  | 
end)  | 
| 
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57944 
diff
changeset
 | 
741  | 
end;  | 
| 
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57944 
diff
changeset
 | 
742  | 
|
| 
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57944 
diff
changeset
 | 
743  | 
fun syntax scan src = apsnd Context.the_proof o syntax_generic scan src o Context.Proof;  | 
| 
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57944 
diff
changeset
 | 
744  | 
|
| 5825 | 745  | 
end;  | 
| 
58011
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57944 
diff
changeset
 | 
746  | 
|
| 
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57944 
diff
changeset
 | 
747  | 
type 'a parser = 'a Token.parser;  | 
| 
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57944 
diff
changeset
 | 
748  | 
type 'a context_parser = 'a Token.context_parser;  |