| author | wenzelm | 
| Fri, 19 Aug 2022 16:46:00 +0200 | |
| changeset 75906 | 2167b9e3157a | 
| parent 74262 | 839a6e284545 | 
| child 77769 | 17391f298cf5 | 
| permissions | -rw-r--r-- | 
| 6118 | 1 | (* Title: Pure/General/position.ML | 
| 74165 | 2 | Author: Makarius | 
| 5010 | 3 | |
| 74165 | 4 | Source positions starting from 1; values <= 0 mean "absent". Count Isabelle | 
| 5 | symbols, not UTF8 bytes nor UTF16 characters. Position range specifies a | |
| 6 | right-open interval offset .. end_offset (exclusive). | |
| 5010 | 7 | *) | 
| 8 | ||
| 9 | signature POSITION = | |
| 10 | sig | |
| 41484 | 11 | eqtype T | 
| 71465 
910a081cca74
more position information for oracles (e.g. "skip_proof" for 'sorry'), requires Proofterm.proofs := 1;
 wenzelm parents: 
70498diff
changeset | 12 | val ord: T ord | 
| 62940 | 13 | val make: Thread_Position.T -> T | 
| 14 | val dest: T -> Thread_Position.T | |
| 23673 | 15 | val line_of: T -> int option | 
| 27795 | 16 | val offset_of: T -> int option | 
| 41484 | 17 | val end_offset_of: T -> int option | 
| 23673 | 18 | val file_of: T -> string option | 
| 74166 | 19 | val id_of: T -> string option | 
| 74174 | 20 | val symbol: Symbol.symbol -> T -> T | 
| 21 | val symbol_explode: string -> T -> T | |
| 68177 | 22 | val distance_of: T * T -> int option | 
| 5010 | 23 | val none: T | 
| 27777 | 24 | val start: T | 
| 31435 | 25 | val file_name: string -> Properties.T | 
| 44200 | 26 | val file_only: string -> T | 
| 27777 | 27 | val file: string -> T | 
| 56437 | 28 | val line_file_only: int -> string -> T | 
| 29 | val line_file: int -> string -> T | |
| 5010 | 30 | val line: int -> T | 
| 67280 
dfc5a1503916
clarified default position for empty message pos;
 wenzelm parents: 
64556diff
changeset | 31 | val get_props: T -> Properties.T | 
| 29307 | 32 | val id: string -> T | 
| 32573 | 33 | val id_only: string -> T | 
| 27426 | 34 | val put_id: string -> T -> T | 
| 68858 | 35 | val copy_id: T -> T -> T | 
| 68829 | 36 | val id_properties_of: T -> Properties.T | 
| 50911 
ee7fe4230642
more explicit treatment of (optional) exception properties, notably for "serial" -- avoid conflict with startPosition = offset;
 wenzelm parents: 
50254diff
changeset | 37 | val parse_id: T -> int option | 
| 74180 | 38 | val advance_offsets: int -> T -> T | 
| 68183 
6560324b1e4d
adjust position according to offset of command/exec id;
 wenzelm parents: 
68177diff
changeset | 39 | val adjust_offsets: (int -> int option) -> T -> T | 
| 28017 | 40 | val of_properties: Properties.T -> T | 
| 41 | val properties_of: T -> Properties.T | |
| 68172 | 42 | val offset_properties_of: T -> Properties.T | 
| 42327 
7c7cc7590eb3
Name_Space.entry_markup: keep def position as separate properties;
 wenzelm parents: 
42204diff
changeset | 43 | val def_properties_of: T -> Properties.T | 
| 71910 | 44 | val entity_markup: string -> string * T -> Markup.T | 
| 74262 | 45 |   val make_entity_markup: {def: bool} -> serial -> string -> string * T -> Markup.T
 | 
| 39440 
4c2547af5909
simplified/clarified (Context_)Position.markup/reported_text;
 wenzelm parents: 
38887diff
changeset | 46 | val markup: T -> Markup.T -> Markup.T | 
| 41504 | 47 | val is_reported: T -> bool | 
| 56459 
38d0b2099743
no report for position singularity, notably for aux. file, especially empty one;
 wenzelm parents: 
56437diff
changeset | 48 | val is_reported_range: T -> bool | 
| 39507 
839873937ddd
tuned signature of (Context_)Position.report variants;
 wenzelm parents: 
39440diff
changeset | 49 | val reported_text: T -> Markup.T -> string -> string | 
| 
839873937ddd
tuned signature of (Context_)Position.report variants;
 wenzelm parents: 
39440diff
changeset | 50 | val report_text: T -> Markup.T -> string -> unit | 
| 
839873937ddd
tuned signature of (Context_)Position.report variants;
 wenzelm parents: 
39440diff
changeset | 51 | val report: T -> Markup.T -> unit | 
| 44736 | 52 | type report = T * Markup.T | 
| 48767 | 53 | type report_text = report * string | 
| 54 | val reports_text: report_text list -> unit | |
| 44736 | 55 | val reports: report list -> unit | 
| 55922 
710bc66f432c
more markup for inner syntax class/type names (notably for completion);
 wenzelm parents: 
55624diff
changeset | 56 | val store_reports: report_text list Unsynchronized.ref -> | 
| 
710bc66f432c
more markup for inner syntax class/type names (notably for completion);
 wenzelm parents: 
55624diff
changeset | 57 |     T list -> ('a -> Markup.T list) -> 'a -> unit
 | 
| 55959 
c3b458435f4f
more decisive commitment to get_free vs. the_const;
 wenzelm parents: 
55922diff
changeset | 58 | val append_reports: report_text list Unsynchronized.ref -> report list -> unit | 
| 71465 
910a081cca74
more position information for oracles (e.g. "skip_proof" for 'sorry'), requires Proofterm.proofs := 1;
 wenzelm parents: 
70498diff
changeset | 59 | val here_strs: T -> string * string | 
| 48992 | 60 | val here: T -> string | 
| 49691 | 61 | val here_list: T list -> string | 
| 27764 | 62 | type range = T * T | 
| 27796 | 63 | val no_range: range | 
| 62800 | 64 | val no_range_position: T -> T | 
| 62797 | 65 | val range_position: range -> T | 
| 66 | val range: T * T -> range | |
| 58978 
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
 wenzelm parents: 
58854diff
changeset | 67 | val range_of_properties: Properties.T -> range | 
| 
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
 wenzelm parents: 
58854diff
changeset | 68 | val properties_of_range: range -> Properties.T | 
| 25817 | 69 | val thread_data: unit -> T | 
| 70 |   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: 
48992diff
changeset | 71 | val default: T -> bool * T | 
| 5010 | 72 | end; | 
| 73 | ||
| 74 | structure Position: POSITION = | |
| 75 | struct | |
| 76 | ||
| 77 | (* datatype position *) | |
| 78 | ||
| 74179 | 79 | type count = int * int * int; | 
| 80 | datatype T = Pos of count * Properties.T; | |
| 41484 | 81 | |
| 73864 | 82 | fun dest2 f = apply2 (fn Pos p => f p); | 
| 83 | ||
| 84 | val ord = | |
| 85 | pointer_eq_ord | |
| 86 | (int_ord o dest2 (#1 o #1) ||| | |
| 87 | int_ord o dest2 (#2 o #1) ||| | |
| 88 | int_ord o dest2 (#3 o #1) ||| | |
| 89 | Properties.ord o dest2 #2); | |
| 71465 
910a081cca74
more position information for oracles (e.g. "skip_proof" for 'sorry'), requires Proofterm.proofs := 1;
 wenzelm parents: 
70498diff
changeset | 90 | |
| 41484 | 91 | fun norm_props (props: Properties.T) = | 
| 45666 | 92 | maps (fn a => the_list (find_first (fn (b, _) => a = b) props)) | 
| 72707 | 93 | [Markup.fileN, Markup.idN]; | 
| 41484 | 94 | |
| 43710 
7270ae921cf2
discontinued odd Position.column -- left-over from attempts at PGIP implementation;
 wenzelm parents: 
42818diff
changeset | 95 | 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: 
42818diff
changeset | 96 | fun dest (Pos ((i, j, k), props)) = {line = i, offset = j, end_offset = k, props = props};
 | 
| 27777 | 97 | |
| 27795 | 98 | fun valid (i: int) = i > 0; | 
| 74179 | 99 | val invalid = not o valid; | 
| 74166 | 100 | fun maybe_valid i = if valid i then SOME i else NONE; | 
| 27795 | 101 | fun if_valid i i' = if valid i then i' else i; | 
| 102 | ||
| 27777 | 103 | |
| 27796 | 104 | (* fields *) | 
| 105 | ||
| 74166 | 106 | fun line_of (Pos ((i, _, _), _)) = maybe_valid i; | 
| 107 | fun offset_of (Pos ((_, j, _), _)) = maybe_valid j; | |
| 108 | fun end_offset_of (Pos ((_, _, k), _)) = maybe_valid k; | |
| 27796 | 109 | |
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49691diff
changeset | 110 | fun file_of (Pos (_, props)) = Properties.get props Markup.fileN; | 
| 74166 | 111 | fun id_of (Pos (_, props)) = Properties.get props Markup.idN; | 
| 27796 | 112 | |
| 113 | ||
| 74176 | 114 | (* count position *) | 
| 26003 | 115 | |
| 74179 | 116 | fun count_symbol "\n" ((i, j, k): count) = | 
| 43710 
7270ae921cf2
discontinued odd Position.column -- left-over from attempts at PGIP implementation;
 wenzelm parents: 
42818diff
changeset | 117 | (if_valid i (i + 1), if_valid j (j + 1), k) | 
| 74176 | 118 | | count_symbol s (i, j, k) = | 
| 119 | if Symbol.not_eof s then (i, if_valid j (j + 1), k) else (i, j, k); | |
| 5010 | 120 | |
| 74179 | 121 | fun count_invalid ((i, j, _): count) = invalid i andalso invalid j; | 
| 27795 | 122 | |
| 74174 | 123 | fun symbol sym (pos as (Pos (count, props))) = | 
| 74179 | 124 | if count_invalid count then pos else Pos (count_symbol sym count, props); | 
| 27777 | 125 | |
| 74174 | 126 | val symbol_explode = fold symbol o Symbol.explode; | 
| 74171 | 127 | |
| 74176 | 128 | |
| 27796 | 129 | (* distance of adjacent positions *) | 
| 27777 | 130 | |
| 68177 | 131 | fun distance_of (Pos ((_, j, _), _), Pos ((_, j', _), _)) = | 
| 132 | if valid j andalso valid j' then SOME (j' - j) else NONE; | |
| 26882 | 133 | |
| 5010 | 134 | |
| 27777 | 135 | (* make position *) | 
| 136 | ||
| 43710 
7270ae921cf2
discontinued odd Position.column -- left-over from attempts at PGIP implementation;
 wenzelm parents: 
42818diff
changeset | 137 | val none = Pos ((0, 0, 0), []); | 
| 
7270ae921cf2
discontinued odd Position.column -- left-over from attempts at PGIP implementation;
 wenzelm parents: 
42818diff
changeset | 138 | val start = Pos ((1, 1, 0), []); | 
| 27744 | 139 | |
| 29307 | 140 | |
| 27796 | 141 | fun file_name "" = [] | 
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49691diff
changeset | 142 | | file_name name = [(Markup.fileN, name)]; | 
| 27796 | 143 | |
| 44200 | 144 | 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: 
42818diff
changeset | 145 | fun file name = Pos ((1, 1, 0), file_name name); | 
| 27744 | 146 | |
| 56437 | 147 | 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: 
42818diff
changeset | 148 | fun line_file i name = Pos ((i, 1, 0), file_name name); | 
| 27795 | 149 | fun line i = line_file i ""; | 
| 5010 | 150 | |
| 67280 
dfc5a1503916
clarified default position for empty message pos;
 wenzelm parents: 
64556diff
changeset | 151 | fun get_props (Pos (_, props)) = props; | 
| 
dfc5a1503916
clarified default position for empty message pos;
 wenzelm parents: 
64556diff
changeset | 152 | |
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49691diff
changeset | 153 | 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: 
49691diff
changeset | 154 | fun id_only id = Pos ((0, 0, 0), [(Markup.idN, id)]); | 
| 22158 | 155 | |
| 62750 | 156 | fun put_id id (Pos (count, props)) = Pos (count, norm_props (Properties.put (Markup.idN, id) props)); | 
| 74166 | 157 | fun copy_id pos = (case id_of pos of NONE => I | SOME id => put_id id); | 
| 27426 | 158 | |
| 74166 | 159 | fun parse_id pos = Option.map Value.parse_int (id_of pos); | 
| 50911 
ee7fe4230642
more explicit treatment of (optional) exception properties, notably for "serial" -- avoid conflict with startPosition = offset;
 wenzelm parents: 
50254diff
changeset | 160 | |
| 68829 | 161 | fun id_properties_of pos = | 
| 74166 | 162 | (case id_of pos of | 
| 68829 | 163 | SOME id => [(Markup.idN, id)] | 
| 164 | | NONE => []); | |
| 29307 | 165 | |
| 74180 | 166 | |
| 167 | (* adjust offsets *) | |
| 168 | ||
| 169 | fun advance_offsets offset (pos as (Pos (count as (i, j, k), props))) = | |
| 170 | if offset = 0 orelse count_invalid count then pos | |
| 171 | else if offset < 0 then raise Fail "Illegal offset" | |
| 172 | else if valid i then raise Fail "Illegal line position" | |
| 173 | else Pos ((i, if_valid j (j + offset), if_valid k (k + offset)), props); | |
| 174 | ||
| 68183 
6560324b1e4d
adjust position according to offset of command/exec id;
 wenzelm parents: 
68177diff
changeset | 175 | fun adjust_offsets adjust (pos as Pos (_, props)) = | 
| 68858 | 176 | if is_none (file_of pos) then | 
| 177 | (case parse_id pos of | |
| 178 | SOME id => | |
| 179 | (case adjust id of | |
| 180 | SOME offset => | |
| 181 | let val Pos (count, _) = advance_offsets offset pos | |
| 182 | in Pos (count, Properties.remove Markup.idN props) end | |
| 183 | | NONE => pos) | |
| 184 | | NONE => pos) | |
| 185 | else pos; | |
| 68183 
6560324b1e4d
adjust position according to offset of command/exec id;
 wenzelm parents: 
68177diff
changeset | 186 | |
| 
6560324b1e4d
adjust position according to offset of command/exec id;
 wenzelm parents: 
68177diff
changeset | 187 | |
| 29307 | 188 | (* markup properties *) | 
| 189 | ||
| 74181 | 190 | fun get_int props name = | 
| 58978 
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
 wenzelm parents: 
58854diff
changeset | 191 | (case Properties.get props name of | 
| 
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
 wenzelm parents: 
58854diff
changeset | 192 | NONE => 0 | 
| 63806 | 193 | | SOME s => Value.parse_int s); | 
| 58978 
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
 wenzelm parents: 
58854diff
changeset | 194 | |
| 25817 | 195 | fun of_properties props = | 
| 74181 | 196 |   make {
 | 
| 197 | line = get_int props Markup.lineN, | |
| 198 | offset = get_int props Markup.offsetN, | |
| 199 | end_offset = get_int props Markup.end_offsetN, | |
| 58978 
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
 wenzelm parents: 
58854diff
changeset | 200 | props = props}; | 
| 26003 | 201 | |
| 74181 | 202 | fun int_entry k i = if invalid i then [] else [(k, Value.print_int i)]; | 
| 41484 | 203 | |
| 43710 
7270ae921cf2
discontinued odd Position.column -- left-over from attempts at PGIP implementation;
 wenzelm parents: 
42818diff
changeset | 204 | fun properties_of (Pos ((i, j, k), props)) = | 
| 74181 | 205 | int_entry Markup.lineN i @ | 
| 206 | int_entry Markup.offsetN j @ | |
| 207 | int_entry Markup.end_offsetN k @ props; | |
| 26003 | 208 | |
| 68172 | 209 | fun offset_properties_of (Pos ((_, j, k), _)) = | 
| 74181 | 210 | int_entry Markup.offsetN j @ | 
| 211 | int_entry Markup.end_offsetN k; | |
| 68172 | 212 | |
| 74182 | 213 | val def_properties_of = properties_of #> map (apfst Markup.def_name); | 
| 42327 
7c7cc7590eb3
Name_Space.entry_markup: keep def position as separate properties;
 wenzelm parents: 
42204diff
changeset | 214 | |
| 71910 | 215 | fun entity_markup kind (name, pos) = | 
| 216 | Markup.entity kind name |> Markup.properties (def_properties_of pos); | |
| 217 | ||
| 74262 | 218 | fun make_entity_markup {def} serial kind (name, pos) =
 | 
| 74183 | 219 | let | 
| 220 | val props = | |
| 221 | if def then (Markup.defN, Value.print_int serial) :: properties_of pos | |
| 222 | else (Markup.refN, Value.print_int serial) :: def_properties_of pos; | |
| 223 | in Markup.entity kind name |> Markup.properties props end; | |
| 45412 | 224 | |
| 39440 
4c2547af5909
simplified/clarified (Context_)Position.markup/reported_text;
 wenzelm parents: 
38887diff
changeset | 225 | val markup = Markup.properties o properties_of; | 
| 
4c2547af5909
simplified/clarified (Context_)Position.markup/reported_text;
 wenzelm parents: 
38887diff
changeset | 226 | |
| 
4c2547af5909
simplified/clarified (Context_)Position.markup/reported_text;
 wenzelm parents: 
38887diff
changeset | 227 | |
| 
4c2547af5909
simplified/clarified (Context_)Position.markup/reported_text;
 wenzelm parents: 
38887diff
changeset | 228 | (* 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: 
38236diff
changeset | 229 | |
| 74166 | 230 | fun is_reported pos = is_some (offset_of pos) andalso is_some (id_of pos); | 
| 56459 
38d0b2099743
no report for position singularity, notably for aux. file, especially empty one;
 wenzelm parents: 
56437diff
changeset | 231 | 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: 
29307diff
changeset | 232 | |
| 41504 | 233 | 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: 
55959diff
changeset | 234 | fun report_text pos markup txt = Output.report [reported_text pos markup txt]; | 
| 39507 
839873937ddd
tuned signature of (Context_)Position.report variants;
 wenzelm parents: 
39440diff
changeset | 235 | fun report pos markup = report_text pos markup ""; | 
| 27764 | 236 | |
| 44736 | 237 | type report = T * Markup.T; | 
| 48767 | 238 | type report_text = report * string; | 
| 44736 | 239 | |
| 48767 | 240 | val reports_text = | 
| 241 | 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: 
55959diff
changeset | 242 | #> Output.report; | 
| 42204 | 243 | |
| 48767 | 244 | val reports = map (rpair "") #> reports_text; | 
| 245 | ||
| 44735 | 246 | fun store_reports _ [] _ _ = () | 
| 55922 
710bc66f432c
more markup for inner syntax class/type names (notably for completion);
 wenzelm parents: 
55624diff
changeset | 247 | | store_reports (r: report_text list Unsynchronized.ref) ps markup x = | 
| 42204 | 248 | let val ms = markup x | 
| 55922 
710bc66f432c
more markup for inner syntax class/type names (notably for completion);
 wenzelm parents: 
55624diff
changeset | 249 | in Unsynchronized.change r (fold (fn p => fold (fn m => cons ((p, m), "")) ms) ps) end; | 
| 42204 | 250 | |
| 55959 
c3b458435f4f
more decisive commitment to get_free vs. the_const;
 wenzelm parents: 
55922diff
changeset | 251 | fun append_reports (r: report_text list Unsynchronized.ref) reports = | 
| 
c3b458435f4f
more decisive commitment to get_free vs. the_const;
 wenzelm parents: 
55922diff
changeset | 252 | Unsynchronized.change r (append (map (rpair "") reports)); | 
| 
c3b458435f4f
more decisive commitment to get_free vs. the_const;
 wenzelm parents: 
55922diff
changeset | 253 | |
| 25817 | 254 | |
| 56532 
3da244bc02bd
tuned message, to accommodate extra brackets produced by Scala parsers;
 wenzelm parents: 
56459diff
changeset | 255 | (* here: user output *) | 
| 26003 | 256 | |
| 71465 
910a081cca74
more position information for oracles (e.g. "skip_proof" for 'sorry'), requires Proofterm.proofs := 1;
 wenzelm parents: 
70498diff
changeset | 257 | fun here_strs pos = | 
| 
910a081cca74
more position information for oracles (e.g. "skip_proof" for 'sorry'), requires Proofterm.proofs := 1;
 wenzelm parents: 
70498diff
changeset | 258 | (case (line_of pos, file_of pos) of | 
| 
910a081cca74
more position information for oracles (e.g. "skip_proof" for 'sorry'), requires Proofterm.proofs := 1;
 wenzelm parents: 
70498diff
changeset | 259 |     (SOME i, NONE) => (" ", "(line " ^ Value.print_int i ^ ")")
 | 
| 
910a081cca74
more position information for oracles (e.g. "skip_proof" for 'sorry'), requires Proofterm.proofs := 1;
 wenzelm parents: 
70498diff
changeset | 260 |   | (SOME i, SOME name) => (" ", "(line " ^ Value.print_int i ^ " of " ^ quote name ^ ")")
 | 
| 
910a081cca74
more position information for oracles (e.g. "skip_proof" for 'sorry'), requires Proofterm.proofs := 1;
 wenzelm parents: 
70498diff
changeset | 261 |   | (NONE, SOME name) => (" ", "(file " ^ quote name ^ ")")
 | 
| 
910a081cca74
more position information for oracles (e.g. "skip_proof" for 'sorry'), requires Proofterm.proofs := 1;
 wenzelm parents: 
70498diff
changeset | 262 |   | _ => if is_reported pos then ("", "\092<^here>") else ("", ""));
 | 
| 
910a081cca74
more position information for oracles (e.g. "skip_proof" for 'sorry'), requires Proofterm.proofs := 1;
 wenzelm parents: 
70498diff
changeset | 263 | |
| 48992 | 264 | fun here pos = | 
| 26003 | 265 | let | 
| 266 | val props = properties_of pos; | |
| 71465 
910a081cca74
more position information for oracles (e.g. "skip_proof" for 'sorry'), requires Proofterm.proofs := 1;
 wenzelm parents: 
70498diff
changeset | 267 | val (s1, s2) = here_strs pos; | 
| 26003 | 268 | in | 
| 69348 | 269 | if s2 = "" then "" | 
| 55624 | 270 | else s1 ^ Markup.markup (Markup.properties props Markup.position) s2 | 
| 26003 | 271 | end; | 
| 25817 | 272 | |
| 70498 | 273 | val here_list = map here #> distinct (op =) #> implode; | 
| 49691 | 274 | |
| 23627 | 275 | |
| 27736 | 276 | (* range *) | 
| 277 | ||
| 278 | type range = T * T; | |
| 279 | ||
| 27796 | 280 | val no_range = (none, none); | 
| 281 | ||
| 62800 | 282 | fun no_range_position (Pos ((i, j, _), props)) = Pos ((i, j, 0), props); | 
| 62797 | 283 | fun range_position (Pos ((i, j, _), props), Pos ((_, j', _), _)) = Pos ((i, j, j'), props); | 
| 62800 | 284 | fun range (pos, pos') = (range_position (pos, pos'), no_range_position pos'); | 
| 27777 | 285 | |
| 58978 
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
 wenzelm parents: 
58854diff
changeset | 286 | fun range_of_properties props = | 
| 
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
 wenzelm parents: 
58854diff
changeset | 287 | let | 
| 
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
 wenzelm parents: 
58854diff
changeset | 288 | val pos = of_properties props; | 
| 
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
 wenzelm parents: 
58854diff
changeset | 289 | val pos' = | 
| 74181 | 290 |       make {line = get_int props Markup.end_lineN,
 | 
| 291 | offset = get_int props Markup.end_offsetN, | |
| 58978 
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
 wenzelm parents: 
58854diff
changeset | 292 | end_offset = 0, | 
| 
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
 wenzelm parents: 
58854diff
changeset | 293 | props = props}; | 
| 
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
 wenzelm parents: 
58854diff
changeset | 294 | in (pos, pos') end; | 
| 
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
 wenzelm parents: 
58854diff
changeset | 295 | |
| 
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
 wenzelm parents: 
58854diff
changeset | 296 | fun properties_of_range (pos, pos') = | 
| 74181 | 297 | properties_of pos @ int_entry Markup.end_lineN (the_default 0 (line_of pos')); | 
| 58978 
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
 wenzelm parents: 
58854diff
changeset | 298 | |
| 27764 | 299 | |
| 300 | (* thread data *) | |
| 301 | ||
| 62940 | 302 | val thread_data = make o Thread_Position.get; | 
| 303 | fun setmp_thread_data pos = Thread_Position.setmp (dest pos); | |
| 27764 | 304 | |
| 33097 
9d501e11084a
maintain position of formal entities via name space;
 wenzelm parents: 
32573diff
changeset | 305 | fun default pos = | 
| 49528 
789b73fcca72
report proper binding positions only -- avoid swamping document model with unspecific information;
 wenzelm parents: 
48992diff
changeset | 306 | if pos = none then (false, thread_data ()) | 
| 
789b73fcca72
report proper binding positions only -- avoid swamping document model with unspecific information;
 wenzelm parents: 
48992diff
changeset | 307 | else (true, pos); | 
| 33097 
9d501e11084a
maintain position of formal entities via name space;
 wenzelm parents: 
32573diff
changeset | 308 | |
| 27764 | 309 | end; |