| author | wenzelm |
| Sun, 13 May 2018 16:51:50 +0200 | |
| changeset 68170 | 7e1daf6f2578 |
| parent 67280 | dfc5a1503916 |
| child 68172 | 0f14cf9c632f |
| permissions | -rw-r--r-- |
| 6118 | 1 |
(* Title: Pure/General/position.ML |
| 5010 | 2 |
Author: Markus Wenzel, TU Muenchen |
3 |
||
| 29307 | 4 |
Source positions: counting Isabelle symbols, starting from 1. |
| 5010 | 5 |
*) |
6 |
||
7 |
signature POSITION = |
|
8 |
sig |
|
| 41484 | 9 |
eqtype T |
| 62940 | 10 |
val make: Thread_Position.T -> T |
11 |
val dest: T -> Thread_Position.T |
|
| 23673 | 12 |
val line_of: T -> int option |
| 27795 | 13 |
val offset_of: T -> int option |
| 41484 | 14 |
val end_offset_of: T -> int option |
| 23673 | 15 |
val file_of: T -> string option |
| 27796 | 16 |
val advance: Symbol.symbol -> T -> T |
|
59085
08a6901eb035
clarified define_command: send tokens more directly, without requiring keywords in ML;
wenzelm
parents:
58978
diff
changeset
|
17 |
val advance_offset: int -> T -> T |
| 27796 | 18 |
val distance_of: T -> T -> int |
| 5010 | 19 |
val none: T |
| 27777 | 20 |
val start: T |
| 31435 | 21 |
val file_name: string -> Properties.T |
| 44200 | 22 |
val file_only: string -> T |
| 27777 | 23 |
val file: string -> T |
| 56437 | 24 |
val line_file_only: int -> string -> T |
25 |
val line_file: int -> string -> T |
|
| 5010 | 26 |
val line: int -> T |
|
67280
dfc5a1503916
clarified default position for empty message pos;
wenzelm
parents:
64556
diff
changeset
|
27 |
val get_props: T -> Properties.T |
| 29307 | 28 |
val id: string -> T |
| 32573 | 29 |
val id_only: string -> T |
| 27426 | 30 |
val get_id: T -> string option |
31 |
val put_id: string -> T -> T |
|
|
50911
ee7fe4230642
more explicit treatment of (optional) exception properties, notably for "serial" -- avoid conflict with startPosition = offset;
wenzelm
parents:
50254
diff
changeset
|
32 |
val parse_id: T -> int option |
| 28017 | 33 |
val of_properties: Properties.T -> T |
34 |
val properties_of: T -> Properties.T |
|
|
42327
7c7cc7590eb3
Name_Space.entry_markup: keep def position as separate properties;
wenzelm
parents:
42204
diff
changeset
|
35 |
val def_properties_of: T -> Properties.T |
| 45412 | 36 |
val entity_properties_of: bool -> serial -> T -> Properties.T |
| 28017 | 37 |
val default_properties: T -> Properties.T -> Properties.T |
|
39440
4c2547af5909
simplified/clarified (Context_)Position.markup/reported_text;
wenzelm
parents:
38887
diff
changeset
|
38 |
val markup: T -> Markup.T -> Markup.T |
| 41504 | 39 |
val is_reported: T -> bool |
|
56459
38d0b2099743
no report for position singularity, notably for aux. file, especially empty one;
wenzelm
parents:
56437
diff
changeset
|
40 |
val is_reported_range: T -> bool |
|
39507
839873937ddd
tuned signature of (Context_)Position.report variants;
wenzelm
parents:
39440
diff
changeset
|
41 |
val reported_text: T -> Markup.T -> string -> string |
|
839873937ddd
tuned signature of (Context_)Position.report variants;
wenzelm
parents:
39440
diff
changeset
|
42 |
val report_text: T -> Markup.T -> string -> unit |
|
839873937ddd
tuned signature of (Context_)Position.report variants;
wenzelm
parents:
39440
diff
changeset
|
43 |
val report: T -> Markup.T -> unit |
| 44736 | 44 |
type report = T * Markup.T |
| 48767 | 45 |
type report_text = report * string |
46 |
val reports_text: report_text list -> unit |
|
| 44736 | 47 |
val reports: report list -> unit |
|
55922
710bc66f432c
more markup for inner syntax class/type names (notably for completion);
wenzelm
parents:
55624
diff
changeset
|
48 |
val store_reports: report_text list Unsynchronized.ref -> |
|
710bc66f432c
more markup for inner syntax class/type names (notably for completion);
wenzelm
parents:
55624
diff
changeset
|
49 |
T list -> ('a -> Markup.T list) -> 'a -> unit
|
|
55959
c3b458435f4f
more decisive commitment to get_free vs. the_const;
wenzelm
parents:
55922
diff
changeset
|
50 |
val append_reports: report_text list Unsynchronized.ref -> report list -> unit |
| 48992 | 51 |
val here: T -> string |
| 49691 | 52 |
val here_list: T list -> string |
| 27764 | 53 |
type range = T * T |
| 27796 | 54 |
val no_range: range |
| 62800 | 55 |
val no_range_position: T -> T |
| 62797 | 56 |
val range_position: range -> T |
57 |
val range: T * T -> range |
|
|
58978
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
wenzelm
parents:
58854
diff
changeset
|
58 |
val range_of_properties: Properties.T -> range |
|
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
wenzelm
parents:
58854
diff
changeset
|
59 |
val properties_of_range: range -> Properties.T |
| 25817 | 60 |
val thread_data: unit -> T |
61 |
val setmp_thread_data: T -> ('a -> 'b) -> 'a -> 'b
|
|
|
49528
789b73fcca72
report proper binding positions only -- avoid swamping document model with unspecific information;
wenzelm
parents:
48992
diff
changeset
|
62 |
val default: T -> bool * T |
| 5010 | 63 |
end; |
64 |
||
65 |
structure Position: POSITION = |
|
66 |
struct |
|
67 |
||
68 |
(* datatype position *) |
|
69 |
||
|
43710
7270ae921cf2
discontinued odd Position.column -- left-over from attempts at PGIP implementation;
wenzelm
parents:
42818
diff
changeset
|
70 |
datatype T = Pos of (int * int * int) * Properties.T; |
| 41484 | 71 |
|
72 |
fun norm_props (props: Properties.T) = |
|
| 45666 | 73 |
maps (fn a => the_list (find_first (fn (b, _) => a = b) props)) |
|
50201
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents:
49691
diff
changeset
|
74 |
Markup.position_properties'; |
| 41484 | 75 |
|
|
43710
7270ae921cf2
discontinued odd Position.column -- left-over from attempts at PGIP implementation;
wenzelm
parents:
42818
diff
changeset
|
76 |
fun make {line = i, offset = j, end_offset = k, props} = Pos ((i, j, k), norm_props props);
|
|
7270ae921cf2
discontinued odd Position.column -- left-over from attempts at PGIP implementation;
wenzelm
parents:
42818
diff
changeset
|
77 |
fun dest (Pos ((i, j, k), props)) = {line = i, offset = j, end_offset = k, props = props};
|
| 27777 | 78 |
|
| 27795 | 79 |
fun valid (i: int) = i > 0; |
80 |
fun if_valid i i' = if valid i then i' else i; |
|
81 |
||
| 27777 | 82 |
|
| 27796 | 83 |
(* fields *) |
84 |
||
|
43710
7270ae921cf2
discontinued odd Position.column -- left-over from attempts at PGIP implementation;
wenzelm
parents:
42818
diff
changeset
|
85 |
fun line_of (Pos ((i, _, _), _)) = if valid i then SOME i else NONE; |
|
7270ae921cf2
discontinued odd Position.column -- left-over from attempts at PGIP implementation;
wenzelm
parents:
42818
diff
changeset
|
86 |
fun offset_of (Pos ((_, j, _), _)) = if valid j then SOME j else NONE; |
|
7270ae921cf2
discontinued odd Position.column -- left-over from attempts at PGIP implementation;
wenzelm
parents:
42818
diff
changeset
|
87 |
fun end_offset_of (Pos ((_, _, k), _)) = if valid k then SOME k else NONE; |
| 27796 | 88 |
|
|
50201
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents:
49691
diff
changeset
|
89 |
fun file_of (Pos (_, props)) = Properties.get props Markup.fileN; |
| 27796 | 90 |
|
91 |
||
| 27777 | 92 |
(* advance *) |
| 26003 | 93 |
|
|
43710
7270ae921cf2
discontinued odd Position.column -- left-over from attempts at PGIP implementation;
wenzelm
parents:
42818
diff
changeset
|
94 |
fun advance_count "\n" (i: int, j: int, k: int) = |
|
7270ae921cf2
discontinued odd Position.column -- left-over from attempts at PGIP implementation;
wenzelm
parents:
42818
diff
changeset
|
95 |
(if_valid i (i + 1), if_valid j (j + 1), k) |
|
7270ae921cf2
discontinued odd Position.column -- left-over from attempts at PGIP implementation;
wenzelm
parents:
42818
diff
changeset
|
96 |
| advance_count s (i, j, k) = |
| 58854 | 97 |
if Symbol.not_eof s then (i, if_valid j (j + 1), k) |
|
43710
7270ae921cf2
discontinued odd Position.column -- left-over from attempts at PGIP implementation;
wenzelm
parents:
42818
diff
changeset
|
98 |
else (i, j, k); |
| 5010 | 99 |
|
|
43710
7270ae921cf2
discontinued odd Position.column -- left-over from attempts at PGIP implementation;
wenzelm
parents:
42818
diff
changeset
|
100 |
fun invalid_count (i, j, _: int) = |
|
7270ae921cf2
discontinued odd Position.column -- left-over from attempts at PGIP implementation;
wenzelm
parents:
42818
diff
changeset
|
101 |
not (valid i orelse valid j); |
| 27795 | 102 |
|
103 |
fun advance sym (pos as (Pos (count, props))) = |
|
104 |
if invalid_count count then pos else Pos (advance_count sym count, props); |
|
| 27777 | 105 |
|
|
59085
08a6901eb035
clarified define_command: send tokens more directly, without requiring keywords in ML;
wenzelm
parents:
58978
diff
changeset
|
106 |
fun advance_offset offset (pos as (Pos (count as (i, j, k), props))) = |
|
08a6901eb035
clarified define_command: send tokens more directly, without requiring keywords in ML;
wenzelm
parents:
58978
diff
changeset
|
107 |
if invalid_count count then pos |
|
08a6901eb035
clarified define_command: send tokens more directly, without requiring keywords in ML;
wenzelm
parents:
58978
diff
changeset
|
108 |
else if valid i then raise Fail "Illegal line position" |
|
08a6901eb035
clarified define_command: send tokens more directly, without requiring keywords in ML;
wenzelm
parents:
58978
diff
changeset
|
109 |
else Pos ((i, if_valid j (j + offset), k), props); |
|
08a6901eb035
clarified define_command: send tokens more directly, without requiring keywords in ML;
wenzelm
parents:
58978
diff
changeset
|
110 |
|
| 27777 | 111 |
|
| 27796 | 112 |
(* distance of adjacent positions *) |
| 27777 | 113 |
|
|
43710
7270ae921cf2
discontinued odd Position.column -- left-over from attempts at PGIP implementation;
wenzelm
parents:
42818
diff
changeset
|
114 |
fun distance_of (Pos ((_, j, _), _)) (Pos ((_, j', _), _)) = |
| 27796 | 115 |
if valid j andalso valid j' then j' - j |
116 |
else 0; |
|
| 26882 | 117 |
|
| 5010 | 118 |
|
| 27777 | 119 |
(* make position *) |
120 |
||
|
43710
7270ae921cf2
discontinued odd Position.column -- left-over from attempts at PGIP implementation;
wenzelm
parents:
42818
diff
changeset
|
121 |
val none = Pos ((0, 0, 0), []); |
|
7270ae921cf2
discontinued odd Position.column -- left-over from attempts at PGIP implementation;
wenzelm
parents:
42818
diff
changeset
|
122 |
val start = Pos ((1, 1, 0), []); |
| 27744 | 123 |
|
| 29307 | 124 |
|
| 27796 | 125 |
fun file_name "" = [] |
|
50201
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents:
49691
diff
changeset
|
126 |
| file_name name = [(Markup.fileN, name)]; |
| 27796 | 127 |
|
| 44200 | 128 |
fun file_only name = Pos ((0, 0, 0), file_name name); |
|
43710
7270ae921cf2
discontinued odd Position.column -- left-over from attempts at PGIP implementation;
wenzelm
parents:
42818
diff
changeset
|
129 |
fun file name = Pos ((1, 1, 0), file_name name); |
| 27744 | 130 |
|
| 56437 | 131 |
fun line_file_only i name = Pos ((i, 0, 0), file_name name); |
|
43710
7270ae921cf2
discontinued odd Position.column -- left-over from attempts at PGIP implementation;
wenzelm
parents:
42818
diff
changeset
|
132 |
fun line_file i name = Pos ((i, 1, 0), file_name name); |
| 27795 | 133 |
fun line i = line_file i ""; |
| 5010 | 134 |
|
|
67280
dfc5a1503916
clarified default position for empty message pos;
wenzelm
parents:
64556
diff
changeset
|
135 |
fun get_props (Pos (_, props)) = props; |
|
dfc5a1503916
clarified default position for empty message pos;
wenzelm
parents:
64556
diff
changeset
|
136 |
|
|
50201
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents:
49691
diff
changeset
|
137 |
fun id id = Pos ((0, 1, 0), [(Markup.idN, id)]); |
|
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents:
49691
diff
changeset
|
138 |
fun id_only id = Pos ((0, 0, 0), [(Markup.idN, id)]); |
| 22158 | 139 |
|
|
50201
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents:
49691
diff
changeset
|
140 |
fun get_id (Pos (_, props)) = Properties.get props Markup.idN; |
| 62750 | 141 |
fun put_id id (Pos (count, props)) = Pos (count, norm_props (Properties.put (Markup.idN, id) props)); |
| 27426 | 142 |
|
| 63806 | 143 |
fun parse_id pos = Option.map Value.parse_int (get_id pos); |
|
50911
ee7fe4230642
more explicit treatment of (optional) exception properties, notably for "serial" -- avoid conflict with startPosition = offset;
wenzelm
parents:
50254
diff
changeset
|
144 |
|
| 29307 | 145 |
|
146 |
(* markup properties *) |
|
147 |
||
|
58978
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
wenzelm
parents:
58854
diff
changeset
|
148 |
fun get props name = |
|
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
wenzelm
parents:
58854
diff
changeset
|
149 |
(case Properties.get props name of |
|
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
wenzelm
parents:
58854
diff
changeset
|
150 |
NONE => 0 |
| 63806 | 151 |
| SOME s => Value.parse_int s); |
|
58978
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
wenzelm
parents:
58854
diff
changeset
|
152 |
|
| 25817 | 153 |
fun of_properties props = |
|
58978
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
wenzelm
parents:
58854
diff
changeset
|
154 |
make {line = get props Markup.lineN,
|
|
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
wenzelm
parents:
58854
diff
changeset
|
155 |
offset = get props Markup.offsetN, |
|
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
wenzelm
parents:
58854
diff
changeset
|
156 |
end_offset = get props Markup.end_offsetN, |
|
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
wenzelm
parents:
58854
diff
changeset
|
157 |
props = props}; |
| 26003 | 158 |
|
| 63806 | 159 |
fun value k i = if valid i then [(k, Value.print_int i)] else []; |
| 41484 | 160 |
|
|
43710
7270ae921cf2
discontinued odd Position.column -- left-over from attempts at PGIP implementation;
wenzelm
parents:
42818
diff
changeset
|
161 |
fun properties_of (Pos ((i, j, k), props)) = |
|
50201
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents:
49691
diff
changeset
|
162 |
value Markup.lineN i @ value Markup.offsetN j @ value Markup.end_offsetN k @ props; |
| 26003 | 163 |
|
| 54038 | 164 |
val def_properties_of = properties_of #> map (fn (x, y) => ("def_" ^ x, y));
|
|
42327
7c7cc7590eb3
Name_Space.entry_markup: keep def position as separate properties;
wenzelm
parents:
42204
diff
changeset
|
165 |
|
|
57899
5867d1306712
clarified signature: entity serial number is not position id;
wenzelm
parents:
56532
diff
changeset
|
166 |
fun entity_properties_of def serial pos = |
| 63806 | 167 |
if def then (Markup.defN, Value.print_int serial) :: properties_of pos |
168 |
else (Markup.refN, Value.print_int serial) :: def_properties_of pos; |
|
| 45412 | 169 |
|
| 26052 | 170 |
fun default_properties default props = |
|
50201
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents:
49691
diff
changeset
|
171 |
if exists (member (op =) Markup.position_properties o #1) props then props |
| 26052 | 172 |
else properties_of default @ props; |
173 |
||
|
39440
4c2547af5909
simplified/clarified (Context_)Position.markup/reported_text;
wenzelm
parents:
38887
diff
changeset
|
174 |
val markup = Markup.properties o properties_of; |
|
4c2547af5909
simplified/clarified (Context_)Position.markup/reported_text;
wenzelm
parents:
38887
diff
changeset
|
175 |
|
|
4c2547af5909
simplified/clarified (Context_)Position.markup/reported_text;
wenzelm
parents:
38887
diff
changeset
|
176 |
|
|
4c2547af5909
simplified/clarified (Context_)Position.markup/reported_text;
wenzelm
parents:
38887
diff
changeset
|
177 |
(* reports *) |
|
38887
1261481ef5e5
Command.State: add reported positions to markup tree, according main message position or Markup.binding/entity/report occurrences in body;
wenzelm
parents:
38236
diff
changeset
|
178 |
|
| 41504 | 179 |
fun is_reported pos = is_some (offset_of pos) andalso is_some (get_id pos); |
|
56459
38d0b2099743
no report for position singularity, notably for aux. file, especially empty one;
wenzelm
parents:
56437
diff
changeset
|
180 |
fun is_reported_range pos = is_reported pos andalso is_some (end_offset_of pos); |
|
30669
6de7ef888aa3
added report_text -- status messages with text body;
wenzelm
parents:
29307
diff
changeset
|
181 |
|
| 41504 | 182 |
fun reported_text pos m txt = if is_reported pos then Markup.markup (markup pos m) txt else ""; |
|
56333
38f1422ef473
support bulk messages consisting of small string segments, which are more healthy to the Poly/ML RTS and might prevent spurious GC crashes such as MTGCProcessMarkPointers::ScanAddressesInObject;
wenzelm
parents:
55959
diff
changeset
|
183 |
fun report_text pos markup txt = Output.report [reported_text pos markup txt]; |
|
39507
839873937ddd
tuned signature of (Context_)Position.report variants;
wenzelm
parents:
39440
diff
changeset
|
184 |
fun report pos markup = report_text pos markup ""; |
| 27764 | 185 |
|
| 44736 | 186 |
type report = T * Markup.T; |
| 48767 | 187 |
type report_text = report * string; |
| 44736 | 188 |
|
| 48767 | 189 |
val reports_text = |
190 |
map (fn ((pos, m), txt) => if is_reported pos then Markup.markup (markup pos m) txt else "") |
|
|
56333
38f1422ef473
support bulk messages consisting of small string segments, which are more healthy to the Poly/ML RTS and might prevent spurious GC crashes such as MTGCProcessMarkPointers::ScanAddressesInObject;
wenzelm
parents:
55959
diff
changeset
|
191 |
#> Output.report; |
| 42204 | 192 |
|
| 48767 | 193 |
val reports = map (rpair "") #> reports_text; |
194 |
||
| 44735 | 195 |
fun store_reports _ [] _ _ = () |
|
55922
710bc66f432c
more markup for inner syntax class/type names (notably for completion);
wenzelm
parents:
55624
diff
changeset
|
196 |
| store_reports (r: report_text list Unsynchronized.ref) ps markup x = |
| 42204 | 197 |
let val ms = markup x |
|
55922
710bc66f432c
more markup for inner syntax class/type names (notably for completion);
wenzelm
parents:
55624
diff
changeset
|
198 |
in Unsynchronized.change r (fold (fn p => fold (fn m => cons ((p, m), "")) ms) ps) end; |
| 42204 | 199 |
|
|
55959
c3b458435f4f
more decisive commitment to get_free vs. the_const;
wenzelm
parents:
55922
diff
changeset
|
200 |
fun append_reports (r: report_text list Unsynchronized.ref) reports = |
|
c3b458435f4f
more decisive commitment to get_free vs. the_const;
wenzelm
parents:
55922
diff
changeset
|
201 |
Unsynchronized.change r (append (map (rpair "") reports)); |
|
c3b458435f4f
more decisive commitment to get_free vs. the_const;
wenzelm
parents:
55922
diff
changeset
|
202 |
|
| 25817 | 203 |
|
|
56532
3da244bc02bd
tuned message, to accommodate extra brackets produced by Scala parsers;
wenzelm
parents:
56459
diff
changeset
|
204 |
(* here: user output *) |
| 26003 | 205 |
|
| 48992 | 206 |
fun here pos = |
| 26003 | 207 |
let |
208 |
val props = properties_of pos; |
|
| 55624 | 209 |
val (s1, s2) = |
| 26003 | 210 |
(case (line_of pos, file_of pos) of |
| 63806 | 211 |
(SOME i, NONE) => (" ", "(line " ^ Value.print_int i ^ ")")
|
212 |
| (SOME i, SOME name) => (" ", "(line " ^ Value.print_int i ^ " of " ^ quote name ^ ")")
|
|
| 55624 | 213 |
| (NONE, SOME name) => (" ", "(file " ^ quote name ^ ")")
|
| 64556 | 214 |
| _ => if is_reported pos then ("", "\092<^here>") else ("", ""));
|
| 26003 | 215 |
in |
216 |
if null props then "" |
|
| 55624 | 217 |
else s1 ^ Markup.markup (Markup.properties props Markup.position) s2 |
| 26003 | 218 |
end; |
| 25817 | 219 |
|
| 62803 | 220 |
val here_list = implode o map here; |
| 49691 | 221 |
|
| 23627 | 222 |
|
| 27736 | 223 |
(* range *) |
224 |
||
225 |
type range = T * T; |
|
226 |
||
| 27796 | 227 |
val no_range = (none, none); |
228 |
||
| 62800 | 229 |
fun no_range_position (Pos ((i, j, _), props)) = Pos ((i, j, 0), props); |
| 62797 | 230 |
fun range_position (Pos ((i, j, _), props), Pos ((_, j', _), _)) = Pos ((i, j, j'), props); |
| 62800 | 231 |
fun range (pos, pos') = (range_position (pos, pos'), no_range_position pos'); |
| 27777 | 232 |
|
|
58978
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
wenzelm
parents:
58854
diff
changeset
|
233 |
fun range_of_properties props = |
|
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
wenzelm
parents:
58854
diff
changeset
|
234 |
let |
|
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
wenzelm
parents:
58854
diff
changeset
|
235 |
val pos = of_properties props; |
|
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
wenzelm
parents:
58854
diff
changeset
|
236 |
val pos' = |
|
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
wenzelm
parents:
58854
diff
changeset
|
237 |
make {line = get props Markup.end_lineN,
|
|
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
wenzelm
parents:
58854
diff
changeset
|
238 |
offset = get props Markup.end_offsetN, |
|
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
wenzelm
parents:
58854
diff
changeset
|
239 |
end_offset = 0, |
|
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
wenzelm
parents:
58854
diff
changeset
|
240 |
props = props}; |
|
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
wenzelm
parents:
58854
diff
changeset
|
241 |
in (pos, pos') end; |
|
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
wenzelm
parents:
58854
diff
changeset
|
242 |
|
|
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
wenzelm
parents:
58854
diff
changeset
|
243 |
fun properties_of_range (pos, pos') = |
|
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
wenzelm
parents:
58854
diff
changeset
|
244 |
properties_of pos @ value Markup.end_lineN (the_default 0 (line_of pos')); |
|
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
wenzelm
parents:
58854
diff
changeset
|
245 |
|
| 27764 | 246 |
|
247 |
(* thread data *) |
|
248 |
||
| 62940 | 249 |
val thread_data = make o Thread_Position.get; |
250 |
fun setmp_thread_data pos = Thread_Position.setmp (dest pos); |
|
| 27764 | 251 |
|
|
33097
9d501e11084a
maintain position of formal entities via name space;
wenzelm
parents:
32573
diff
changeset
|
252 |
fun default pos = |
|
49528
789b73fcca72
report proper binding positions only -- avoid swamping document model with unspecific information;
wenzelm
parents:
48992
diff
changeset
|
253 |
if pos = none then (false, thread_data ()) |
|
789b73fcca72
report proper binding positions only -- avoid swamping document model with unspecific information;
wenzelm
parents:
48992
diff
changeset
|
254 |
else (true, pos); |
|
33097
9d501e11084a
maintain position of formal entities via name space;
wenzelm
parents:
32573
diff
changeset
|
255 |
|
| 27764 | 256 |
end; |