| author | wenzelm | 
| Sat, 01 Apr 2023 13:58:45 +0200 | |
| changeset 77769 | 17391f298cf5 | 
| parent 74262 | 839a6e284545 | 
| child 77770 | 472e48ed9c79 | 
| 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: 
70498 
diff
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: 
64556 
diff
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: 
50254 
diff
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: 
68177 
diff
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: 
42204 
diff
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: 
38887 
diff
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: 
56437 
diff
changeset
 | 
48  | 
val is_reported_range: T -> bool  | 
| 
39507
 
839873937ddd
tuned signature of (Context_)Position.report variants;
 
wenzelm 
parents: 
39440 
diff
changeset
 | 
49  | 
val reported_text: T -> Markup.T -> string -> string  | 
| 
 
839873937ddd
tuned signature of (Context_)Position.report variants;
 
wenzelm 
parents: 
39440 
diff
changeset
 | 
50  | 
val report_text: T -> Markup.T -> string -> unit  | 
| 
 
839873937ddd
tuned signature of (Context_)Position.report variants;
 
wenzelm 
parents: 
39440 
diff
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: 
55624 
diff
changeset
 | 
56  | 
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
 | 
57  | 
    T list -> ('a -> Markup.T list) -> 'a -> unit
 | 
| 
55959
 
c3b458435f4f
more decisive commitment to get_free vs. the_const;
 
wenzelm 
parents: 
55922 
diff
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: 
70498 
diff
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: 
58854 
diff
changeset
 | 
67  | 
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
 | 
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: 
48992 
diff
changeset
 | 
71  | 
val default: T -> bool * T  | 
| 5010 | 72  | 
end;  | 
73  | 
||
74  | 
structure Position: POSITION =  | 
|
75  | 
struct  | 
|
76  | 
||
77  | 
(* datatype position *)  | 
|
78  | 
||
| 
77769
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
79  | 
datatype T = Pos of {line: int, offset: int, end_offset: int, props: Properties.T};
 | 
| 
71465
 
910a081cca74
more position information for oracles (e.g. "skip_proof" for 'sorry'), requires Proofterm.proofs := 1;
 
wenzelm 
parents: 
70498 
diff
changeset
 | 
80  | 
|
| 41484 | 81  | 
fun norm_props (props: Properties.T) =  | 
| 45666 | 82  | 
maps (fn a => the_list (find_first (fn (b, _) => a = b) props))  | 
| 72707 | 83  | 
[Markup.fileN, Markup.idN];  | 
| 41484 | 84  | 
|
| 
77769
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
85  | 
fun make {line, offset, end_offset, props} =
 | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
86  | 
  Pos {line = line, offset = offset, end_offset = end_offset, props = norm_props props};
 | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
87  | 
|
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
88  | 
fun dest (Pos args) = args;  | 
| 27777 | 89  | 
|
| 
77769
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
90  | 
val ord =  | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
91  | 
pointer_eq_ord  | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
92  | 
(int_ord o apply2 (#line o dest) |||  | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
93  | 
int_ord o apply2 (#offset o dest) |||  | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
94  | 
int_ord o apply2 (#end_offset o dest) |||  | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
95  | 
Properties.ord o apply2 (#props o dest));  | 
| 27795 | 96  | 
|
| 27777 | 97  | 
|
| 27796 | 98  | 
(* fields *)  | 
99  | 
||
| 
77769
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
100  | 
fun valid (i: int) = i > 0;  | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
101  | 
fun maybe_valid i = if valid i then SOME i else NONE;  | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
102  | 
|
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
103  | 
val invalid = not o valid;  | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
104  | 
fun invalid_pos (Pos {line, offset, ...}) = invalid line andalso invalid offset;
 | 
| 27796 | 105  | 
|
| 
77769
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
106  | 
val line_of = maybe_valid o #line o dest;  | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
107  | 
val offset_of = maybe_valid o #offset o dest;  | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
108  | 
val end_offset_of = maybe_valid o #end_offset o dest;  | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
109  | 
|
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
110  | 
fun file_of (Pos {props, ...}) = Properties.get props Markup.fileN;
 | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
111  | 
fun id_of (Pos {props, ...}) = Properties.get props Markup.idN;
 | 
| 27796 | 112  | 
|
113  | 
||
| 
77769
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
114  | 
(* symbol positions *)  | 
| 5010 | 115  | 
|
| 
77769
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
116  | 
fun symbols [] pos = pos  | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
117  | 
  | symbols ss (pos as Pos {line, offset, end_offset, props}) =
 | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
118  | 
let  | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
119  | 
val line' = if valid line then fold (fn s => s = "\n" ? Integer.add 1) ss line else line;  | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
120  | 
val offset' =  | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
121  | 
if valid offset then fold (fn s => Symbol.not_eof s ? Integer.add 1) ss offset else offset;  | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
122  | 
in  | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
123  | 
if line = line' andalso offset = offset' then pos  | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
124  | 
        else Pos {line = line', offset = offset', end_offset = end_offset, props = props}
 | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
125  | 
end;  | 
| 27795 | 126  | 
|
| 
77769
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
127  | 
val symbol = symbols o single;  | 
| 27777 | 128  | 
|
| 
77769
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
129  | 
fun symbol_explode str pos =  | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
130  | 
if str = "" orelse invalid_pos pos then pos  | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
131  | 
else fold symbol (Symbol.explode str) pos;  | 
| 74171 | 132  | 
|
| 74176 | 133  | 
|
| 27796 | 134  | 
(* distance of adjacent positions *)  | 
| 27777 | 135  | 
|
| 
77769
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
136  | 
fun distance_of (Pos {offset, ...}, Pos {offset = offset', ...}) =
 | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
137  | 
if valid offset andalso valid offset' then SOME (offset' - offset) else NONE;  | 
| 26882 | 138  | 
|
| 5010 | 139  | 
|
| 27777 | 140  | 
(* make position *)  | 
141  | 
||
| 
77769
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
142  | 
val none = Pos {line = 0, offset = 0, end_offset = 0, props = []};
 | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
143  | 
val start = Pos {line = 1, offset = 1, end_offset = 0, props = []};
 | 
| 27744 | 144  | 
|
| 29307 | 145  | 
|
| 27796 | 146  | 
fun file_name "" = []  | 
| 
50201
 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 
wenzelm 
parents: 
49691 
diff
changeset
 | 
147  | 
| file_name name = [(Markup.fileN, name)];  | 
| 27796 | 148  | 
|
| 
77769
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
149  | 
fun file_only name = Pos {line = 0, offset = 0, end_offset = 0, props = file_name name};
 | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
150  | 
fun file name = Pos {line = 1, offset = 1, end_offset = 0, props = file_name name};
 | 
| 27744 | 151  | 
|
| 
77769
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
152  | 
fun line_file_only line name = Pos {line = line, offset = 0, end_offset = 0, props = file_name name};
 | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
153  | 
fun line_file line name = Pos {line = line, offset = 1, end_offset = 0, props = file_name name};
 | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
154  | 
fun line line = line_file line "";  | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
155  | 
|
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
156  | 
val get_props = #props o dest;  | 
| 5010 | 157  | 
|
| 
77769
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
158  | 
fun id id = Pos {line = 0, offset = 1, end_offset = 0, props = [(Markup.idN, id)]};
 | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
159  | 
fun id_only id = Pos {line = 0, offset = 0, end_offset = 0, props = [(Markup.idN, id)]};
 | 
| 
67280
 
dfc5a1503916
clarified default position for empty message pos;
 
wenzelm 
parents: 
64556 
diff
changeset
 | 
160  | 
|
| 
77769
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
161  | 
fun put_id id (pos as Pos {line, offset, end_offset, props}) =
 | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
162  | 
let val props' = norm_props (Properties.put (Markup.idN, id) props) in  | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
163  | 
if props = props' then pos  | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
164  | 
    else Pos {line = line, offset = offset, end_offset = end_offset, props = props'}
 | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
165  | 
end;  | 
| 22158 | 166  | 
|
| 74166 | 167  | 
fun copy_id pos = (case id_of pos of NONE => I | SOME id => put_id id);  | 
| 27426 | 168  | 
|
| 74166 | 169  | 
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: 
50254 
diff
changeset
 | 
170  | 
|
| 68829 | 171  | 
fun id_properties_of pos =  | 
| 74166 | 172  | 
(case id_of pos of  | 
| 68829 | 173  | 
SOME id => [(Markup.idN, id)]  | 
174  | 
| NONE => []);  | 
|
| 29307 | 175  | 
|
| 74180 | 176  | 
|
177  | 
(* adjust offsets *)  | 
|
178  | 
||
| 
77769
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
179  | 
fun shift_offsets {remove_id} shift (pos as Pos {line, offset, end_offset, props}) =
 | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
180  | 
if shift < 0 then raise Fail "Illegal offset shift"  | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
181  | 
else if shift > 0 andalso valid line then raise Fail "Illegal line position"  | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
182  | 
else  | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
183  | 
let  | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
184  | 
val offset' = if valid offset then offset + shift else offset;  | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
185  | 
val end_offset' = if valid end_offset then end_offset + shift else end_offset;  | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
186  | 
val props' = if remove_id then Properties.remove Markup.idN props else props;  | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
187  | 
in  | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
188  | 
if offset = offset' andalso end_offset = end_offset' andalso props = props' then pos  | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
189  | 
      else Pos {line = line, offset = offset', end_offset = end_offset', props = props'}
 | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
190  | 
end;  | 
| 74180 | 191  | 
|
| 
77769
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
192  | 
val advance_offsets = shift_offsets {remove_id = false};
 | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
193  | 
|
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
194  | 
fun adjust_offsets adjust pos =  | 
| 68858 | 195  | 
if is_none (file_of pos) then  | 
196  | 
(case parse_id pos of  | 
|
197  | 
SOME id =>  | 
|
198  | 
(case adjust id of  | 
|
| 
77769
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
199  | 
          SOME shift => shift_offsets {remove_id = true} shift pos
 | 
| 68858 | 200  | 
| NONE => pos)  | 
201  | 
| NONE => pos)  | 
|
202  | 
else pos;  | 
|
| 
68183
 
6560324b1e4d
adjust position according to offset of command/exec id;
 
wenzelm 
parents: 
68177 
diff
changeset
 | 
203  | 
|
| 
 
6560324b1e4d
adjust position according to offset of command/exec id;
 
wenzelm 
parents: 
68177 
diff
changeset
 | 
204  | 
|
| 29307 | 205  | 
(* markup properties *)  | 
206  | 
||
| 74181 | 207  | 
fun get_int props name =  | 
| 
58978
 
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
 
wenzelm 
parents: 
58854 
diff
changeset
 | 
208  | 
(case Properties.get props name of  | 
| 
 
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
 
wenzelm 
parents: 
58854 
diff
changeset
 | 
209  | 
NONE => 0  | 
| 63806 | 210  | 
| 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
 | 
211  | 
|
| 25817 | 212  | 
fun of_properties props =  | 
| 74181 | 213  | 
  make {
 | 
214  | 
line = get_int props Markup.lineN,  | 
|
215  | 
offset = get_int props Markup.offsetN,  | 
|
216  | 
end_offset = get_int props Markup.end_offsetN,  | 
|
| 
58978
 
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
 
wenzelm 
parents: 
58854 
diff
changeset
 | 
217  | 
props = props};  | 
| 26003 | 218  | 
|
| 74181 | 219  | 
fun int_entry k i = if invalid i then [] else [(k, Value.print_int i)];  | 
| 41484 | 220  | 
|
| 
77769
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
221  | 
fun properties_of (Pos {line, offset, end_offset, props}) =
 | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
222  | 
int_entry Markup.lineN line @  | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
223  | 
int_entry Markup.offsetN offset @  | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
224  | 
int_entry Markup.end_offsetN end_offset @ props;  | 
| 26003 | 225  | 
|
| 
77769
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
226  | 
fun offset_properties_of (Pos {offset, end_offset, ...}) =
 | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
227  | 
int_entry Markup.offsetN offset @  | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
228  | 
int_entry Markup.end_offsetN end_offset;  | 
| 68172 | 229  | 
|
| 74182 | 230  | 
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: 
42204 
diff
changeset
 | 
231  | 
|
| 71910 | 232  | 
fun entity_markup kind (name, pos) =  | 
233  | 
Markup.entity kind name |> Markup.properties (def_properties_of pos);  | 
|
234  | 
||
| 74262 | 235  | 
fun make_entity_markup {def} serial kind (name, pos) =
 | 
| 74183 | 236  | 
let  | 
237  | 
val props =  | 
|
238  | 
if def then (Markup.defN, Value.print_int serial) :: properties_of pos  | 
|
239  | 
else (Markup.refN, Value.print_int serial) :: def_properties_of pos;  | 
|
240  | 
in Markup.entity kind name |> Markup.properties props end;  | 
|
| 45412 | 241  | 
|
| 
39440
 
4c2547af5909
simplified/clarified (Context_)Position.markup/reported_text;
 
wenzelm 
parents: 
38887 
diff
changeset
 | 
242  | 
val markup = Markup.properties o properties_of;  | 
| 
 
4c2547af5909
simplified/clarified (Context_)Position.markup/reported_text;
 
wenzelm 
parents: 
38887 
diff
changeset
 | 
243  | 
|
| 
 
4c2547af5909
simplified/clarified (Context_)Position.markup/reported_text;
 
wenzelm 
parents: 
38887 
diff
changeset
 | 
244  | 
|
| 
 
4c2547af5909
simplified/clarified (Context_)Position.markup/reported_text;
 
wenzelm 
parents: 
38887 
diff
changeset
 | 
245  | 
(* 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
 | 
246  | 
|
| 74166 | 247  | 
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: 
56437 
diff
changeset
 | 
248  | 
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
 | 
249  | 
|
| 41504 | 250  | 
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
 | 
251  | 
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
 | 
252  | 
fun report pos markup = report_text pos markup "";  | 
| 27764 | 253  | 
|
| 44736 | 254  | 
type report = T * Markup.T;  | 
| 48767 | 255  | 
type report_text = report * string;  | 
| 44736 | 256  | 
|
| 48767 | 257  | 
val reports_text =  | 
258  | 
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
 | 
259  | 
#> Output.report;  | 
| 42204 | 260  | 
|
| 48767 | 261  | 
val reports = map (rpair "") #> reports_text;  | 
262  | 
||
| 44735 | 263  | 
fun store_reports _ [] _ _ = ()  | 
| 
55922
 
710bc66f432c
more markup for inner syntax class/type names (notably for completion);
 
wenzelm 
parents: 
55624 
diff
changeset
 | 
264  | 
| store_reports (r: report_text list Unsynchronized.ref) ps markup x =  | 
| 42204 | 265  | 
let val ms = markup x  | 
| 
55922
 
710bc66f432c
more markup for inner syntax class/type names (notably for completion);
 
wenzelm 
parents: 
55624 
diff
changeset
 | 
266  | 
in Unsynchronized.change r (fold (fn p => fold (fn m => cons ((p, m), "")) ms) ps) end;  | 
| 42204 | 267  | 
|
| 
55959
 
c3b458435f4f
more decisive commitment to get_free vs. the_const;
 
wenzelm 
parents: 
55922 
diff
changeset
 | 
268  | 
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
 | 
269  | 
Unsynchronized.change r (append (map (rpair "") reports));  | 
| 
 
c3b458435f4f
more decisive commitment to get_free vs. the_const;
 
wenzelm 
parents: 
55922 
diff
changeset
 | 
270  | 
|
| 25817 | 271  | 
|
| 
56532
 
3da244bc02bd
tuned message, to accommodate extra brackets produced by Scala parsers;
 
wenzelm 
parents: 
56459 
diff
changeset
 | 
272  | 
(* here: user output *)  | 
| 26003 | 273  | 
|
| 
71465
 
910a081cca74
more position information for oracles (e.g. "skip_proof" for 'sorry'), requires Proofterm.proofs := 1;
 
wenzelm 
parents: 
70498 
diff
changeset
 | 
274  | 
fun here_strs pos =  | 
| 
 
910a081cca74
more position information for oracles (e.g. "skip_proof" for 'sorry'), requires Proofterm.proofs := 1;
 
wenzelm 
parents: 
70498 
diff
changeset
 | 
275  | 
(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: 
70498 
diff
changeset
 | 
276  | 
    (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: 
70498 
diff
changeset
 | 
277  | 
  | (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: 
70498 
diff
changeset
 | 
278  | 
  | (NONE, SOME name) => (" ", "(file " ^ quote name ^ ")")
 | 
| 
 
910a081cca74
more position information for oracles (e.g. "skip_proof" for 'sorry'), requires Proofterm.proofs := 1;
 
wenzelm 
parents: 
70498 
diff
changeset
 | 
279  | 
  | _ => 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: 
70498 
diff
changeset
 | 
280  | 
|
| 48992 | 281  | 
fun here pos =  | 
| 26003 | 282  | 
let  | 
283  | 
val props = properties_of pos;  | 
|
| 
71465
 
910a081cca74
more position information for oracles (e.g. "skip_proof" for 'sorry'), requires Proofterm.proofs := 1;
 
wenzelm 
parents: 
70498 
diff
changeset
 | 
284  | 
val (s1, s2) = here_strs pos;  | 
| 26003 | 285  | 
in  | 
| 69348 | 286  | 
if s2 = "" then ""  | 
| 55624 | 287  | 
else s1 ^ Markup.markup (Markup.properties props Markup.position) s2  | 
| 26003 | 288  | 
end;  | 
| 25817 | 289  | 
|
| 70498 | 290  | 
val here_list = map here #> distinct (op =) #> implode;  | 
| 49691 | 291  | 
|
| 23627 | 292  | 
|
| 27736 | 293  | 
(* range *)  | 
294  | 
||
295  | 
type range = T * T;  | 
|
296  | 
||
| 27796 | 297  | 
val no_range = (none, none);  | 
298  | 
||
| 
77769
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
299  | 
fun no_range_position (Pos {line, offset, end_offset = _, props}) =
 | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
300  | 
  Pos {line = line, offset = offset, end_offset = 0, props = props};
 | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
301  | 
|
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
302  | 
fun range_position (Pos {line, offset, end_offset = _, props}, Pos {offset = offset', ...}) =
 | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
303  | 
  Pos {line = line, offset = offset, end_offset = offset', props = props};
 | 
| 
 
17391f298cf5
misc tuning and clarification: more tight representation;
 
wenzelm 
parents: 
74262 
diff
changeset
 | 
304  | 
|
| 62800 | 305  | 
fun range (pos, pos') = (range_position (pos, pos'), no_range_position pos');  | 
| 27777 | 306  | 
|
| 
58978
 
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
 
wenzelm 
parents: 
58854 
diff
changeset
 | 
307  | 
fun range_of_properties props =  | 
| 
 
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
 
wenzelm 
parents: 
58854 
diff
changeset
 | 
308  | 
let  | 
| 
 
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
 
wenzelm 
parents: 
58854 
diff
changeset
 | 
309  | 
val pos = of_properties props;  | 
| 
 
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
 
wenzelm 
parents: 
58854 
diff
changeset
 | 
310  | 
val pos' =  | 
| 74181 | 311  | 
      make {line = get_int props Markup.end_lineN,
 | 
312  | 
offset = get_int props Markup.end_offsetN,  | 
|
| 
58978
 
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
 
wenzelm 
parents: 
58854 
diff
changeset
 | 
313  | 
end_offset = 0,  | 
| 
 
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
 
wenzelm 
parents: 
58854 
diff
changeset
 | 
314  | 
props = props};  | 
| 
 
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
 
wenzelm 
parents: 
58854 
diff
changeset
 | 
315  | 
in (pos, pos') end;  | 
| 
 
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
 
wenzelm 
parents: 
58854 
diff
changeset
 | 
316  | 
|
| 
 
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
 
wenzelm 
parents: 
58854 
diff
changeset
 | 
317  | 
fun properties_of_range (pos, pos') =  | 
| 74181 | 318  | 
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: 
58854 
diff
changeset
 | 
319  | 
|
| 27764 | 320  | 
|
321  | 
(* thread data *)  | 
|
322  | 
||
| 62940 | 323  | 
val thread_data = make o Thread_Position.get;  | 
324  | 
fun setmp_thread_data pos = Thread_Position.setmp (dest pos);  | 
|
| 27764 | 325  | 
|
| 
33097
 
9d501e11084a
maintain position of formal entities via name space;
 
wenzelm 
parents: 
32573 
diff
changeset
 | 
326  | 
fun default pos =  | 
| 
49528
 
789b73fcca72
report proper binding positions only -- avoid swamping document model with unspecific information;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
327  | 
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
 | 
328  | 
else (true, pos);  | 
| 
33097
 
9d501e11084a
maintain position of formal entities via name space;
 
wenzelm 
parents: 
32573 
diff
changeset
 | 
329  | 
|
| 27764 | 330  | 
end;  |