author | wenzelm |
Sat, 07 Dec 2024 23:50:18 +0100 | |
changeset 81558 | b57996a0688c |
parent 81552 | 4717d3bf5752 |
child 81565 | bf19ea589f99 |
permissions | -rw-r--r-- |
42243 | 1 |
(* Title: Pure/Syntax/syntax_phases.ML |
42241
dd8029f71e1c
separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff
changeset
|
2 |
Author: Makarius |
dd8029f71e1c
separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff
changeset
|
3 |
|
42243 | 4 |
Main phases of inner syntax processing, with standard implementations |
5 |
of parse/unparse operations. |
|
42241
dd8029f71e1c
separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff
changeset
|
6 |
*) |
dd8029f71e1c
separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff
changeset
|
7 |
|
42243 | 8 |
signature SYNTAX_PHASES = |
42242
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
9 |
sig |
74571 | 10 |
val markup_free: Proof.context -> string -> Markup.T list |
63395 | 11 |
val reports_of_scope: Position.T list -> Position.report list |
45427
fca432074fb2
sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents:
45423
diff
changeset
|
12 |
val decode_sort: term -> sort |
fca432074fb2
sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents:
45423
diff
changeset
|
13 |
val decode_typ: term -> typ |
42242
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
14 |
val decode_term: Proof.context -> |
55922
710bc66f432c
more markup for inner syntax class/type names (notably for completion);
wenzelm
parents:
55829
diff
changeset
|
15 |
Position.report_text list * term Exn.result -> Position.report_text list * term Exn.result |
42242
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
16 |
val parse_ast_pattern: Proof.context -> string * string -> Ast.ast |
42247
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
wenzelm
parents:
42245
diff
changeset
|
17 |
val term_of_typ: Proof.context -> typ -> term |
45429 | 18 |
val print_checks: Proof.context -> unit |
19 |
val typ_check: int -> string -> (Proof.context -> typ list -> typ list) -> |
|
20 |
Context.generic -> Context.generic |
|
21 |
val term_check: int -> string -> (Proof.context -> term list -> term list) -> |
|
22 |
Context.generic -> Context.generic |
|
23 |
val typ_uncheck: int -> string -> (Proof.context -> typ list -> typ list) -> |
|
24 |
Context.generic -> Context.generic |
|
25 |
val term_uncheck: int -> string -> (Proof.context -> term list -> term list) -> |
|
26 |
Context.generic -> Context.generic |
|
45444 | 27 |
val typ_check': int -> string -> |
28 |
(typ list -> Proof.context -> (typ list * Proof.context) option) -> |
|
29 |
Context.generic -> Context.generic |
|
30 |
val term_check': int -> string -> |
|
31 |
(term list -> Proof.context -> (term list * Proof.context) option) -> |
|
32 |
Context.generic -> Context.generic |
|
33 |
val typ_uncheck': int -> string -> |
|
34 |
(typ list -> Proof.context -> (typ list * Proof.context) option) -> |
|
35 |
Context.generic -> Context.generic |
|
36 |
val term_uncheck': int -> string -> |
|
37 |
(term list -> Proof.context -> (term list * Proof.context) option) -> |
|
38 |
Context.generic -> Context.generic |
|
42242
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
39 |
end |
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
40 |
|
42243 | 41 |
structure Syntax_Phases: SYNTAX_PHASES = |
42241
dd8029f71e1c
separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff
changeset
|
42 |
struct |
dd8029f71e1c
separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff
changeset
|
43 |
|
42298
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
wenzelm
parents:
42296
diff
changeset
|
44 |
(** markup logical entities **) |
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
wenzelm
parents:
42296
diff
changeset
|
45 |
|
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
wenzelm
parents:
42296
diff
changeset
|
46 |
fun markup_class ctxt c = |
42379 | 47 |
[Name_Space.markup (Type.class_space (Proof_Context.tsig_of ctxt)) c]; |
42298
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
wenzelm
parents:
42296
diff
changeset
|
48 |
|
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
wenzelm
parents:
42296
diff
changeset
|
49 |
fun markup_type ctxt c = |
42379 | 50 |
[Name_Space.markup (Type.type_space (Proof_Context.tsig_of ctxt)) c]; |
42298
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
wenzelm
parents:
42296
diff
changeset
|
51 |
|
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
wenzelm
parents:
42296
diff
changeset
|
52 |
fun markup_const ctxt c = |
42379 | 53 |
[Name_Space.markup (Consts.space_of (Proof_Context.consts_of ctxt)) c]; |
42298
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
wenzelm
parents:
42296
diff
changeset
|
54 |
|
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
wenzelm
parents:
42296
diff
changeset
|
55 |
fun markup_free ctxt x = |
80746
fecfbcf41473
clarified markup: more uniform treatment of parse/print phase;
wenzelm
parents:
80745
diff
changeset
|
56 |
let |
81265
2daa7b2ef46e
clarified (again): Markup.intensify is already part of Variable.markup_fixed for undeclared variable, Markup.fixed is already part of Mariable.markup;
wenzelm
parents:
81260
diff
changeset
|
57 |
val strict = Variable.is_body ctxt andalso not (Syntax.is_pretty_global ctxt); |
2daa7b2ef46e
clarified (again): Markup.intensify is already part of Variable.markup_fixed for undeclared variable, Markup.fixed is already part of Mariable.markup;
wenzelm
parents:
81260
diff
changeset
|
58 |
val m1 = if Variable.is_fixed ctxt x orelse strict then [Variable.markup_fixed ctxt x] else []; |
80746
fecfbcf41473
clarified markup: more uniform treatment of parse/print phase;
wenzelm
parents:
80745
diff
changeset
|
59 |
val m2 = Variable.markup ctxt x; |
81265
2daa7b2ef46e
clarified (again): Markup.intensify is already part of Variable.markup_fixed for undeclared variable, Markup.fixed is already part of Mariable.markup;
wenzelm
parents:
81260
diff
changeset
|
60 |
in m1 @ [m2] end; |
42298
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
wenzelm
parents:
42296
diff
changeset
|
61 |
|
50201
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents:
49690
diff
changeset
|
62 |
fun markup_var xi = [Markup.name (Term.string_of_vname xi) Markup.var]; |
42298
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
wenzelm
parents:
42296
diff
changeset
|
63 |
|
81558
b57996a0688c
clarified term positions and markup: syntax = true means this is via concrete syntax;
wenzelm
parents:
81552
diff
changeset
|
64 |
fun markup_bound def (ps: Term_Position.T list) (name, id) = |
b57996a0688c
clarified term positions and markup: syntax = true means this is via concrete syntax;
wenzelm
parents:
81552
diff
changeset
|
65 |
Markup.bound :: |
b57996a0688c
clarified term positions and markup: syntax = true means this is via concrete syntax;
wenzelm
parents:
81552
diff
changeset
|
66 |
map (fn {pos, ...} => Position.make_entity_markup def id Markup.boundN (name, pos)) ps; |
42298
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
wenzelm
parents:
42296
diff
changeset
|
67 |
|
81552 | 68 |
fun markup_entity ctxt c = |
69 |
Syntax.get_consts (Proof_Context.syntax_of ctxt) c |
|
70 |
|> maps (Lexicon.unmark |
|
71 |
{case_class = markup_class ctxt, |
|
72 |
case_type = markup_type ctxt, |
|
73 |
case_const = markup_const ctxt, |
|
74 |
case_fixed = markup_free ctxt, |
|
75 |
case_default = K []}); |
|
42298
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
wenzelm
parents:
42296
diff
changeset
|
76 |
|
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
wenzelm
parents:
42296
diff
changeset
|
77 |
|
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
wenzelm
parents:
42296
diff
changeset
|
78 |
|
63395 | 79 |
(** reports of implicit variable scope **) |
80 |
||
81 |
fun reports_of_scope [] = [] |
|
82 |
| reports_of_scope (def_pos :: ps) = |
|
83 |
let |
|
84 |
val id = serial (); |
|
74183 | 85 |
fun entity def = Position.make_entity_markup def id "" ("", def_pos); |
63395 | 86 |
in |
87 |
map (rpair Markup.bound) (def_pos :: ps) @ |
|
74262 | 88 |
((def_pos, entity {def = true}) :: map (rpair (entity {def = false})) ps) |
63395 | 89 |
end; |
90 |
||
91 |
||
92 |
||
42242
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
93 |
(** decode parse trees **) |
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
94 |
|
45427
fca432074fb2
sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents:
45423
diff
changeset
|
95 |
(* decode_sort *) |
42242
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
96 |
|
45427
fca432074fb2
sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents:
45423
diff
changeset
|
97 |
fun decode_sort tm = |
42242
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
98 |
let |
45427
fca432074fb2
sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents:
45423
diff
changeset
|
99 |
fun err () = raise TERM ("decode_sort: bad encoding of classes", [tm]); |
42242
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
100 |
|
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
101 |
fun class s = Lexicon.unmark_class s handle Fail _ => err (); |
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
102 |
|
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
103 |
fun classes (Const (s, _)) = [class s] |
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
104 |
| classes (Const ("_classes", _) $ Const (s, _) $ cs) = class s :: classes cs |
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
105 |
| classes _ = err (); |
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
106 |
|
67718 | 107 |
fun sort (Const ("_dummy_sort", _)) = dummyS |
108 |
| sort (Const ("_topsort", _)) = [] |
|
49674
dbadb4d03cbc
report sort assignment of visible type variables;
wenzelm
parents:
49662
diff
changeset
|
109 |
| sort (Const ("_sort", _) $ cs) = classes cs |
42242
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
110 |
| sort (Const (s, _)) = [class s] |
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
111 |
| sort _ = err (); |
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
112 |
in sort tm end; |
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
113 |
|
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
114 |
|
45427
fca432074fb2
sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents:
45423
diff
changeset
|
115 |
(* decode_typ *) |
42242
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
116 |
|
81232 | 117 |
fun decode_pos (Free (x, _)) = if Term_Position.detect x then SOME x else NONE |
49685
4341e4d86872
allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents:
49674
diff
changeset
|
118 |
| decode_pos _ = NONE; |
4341e4d86872
allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents:
49674
diff
changeset
|
119 |
|
45427
fca432074fb2
sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents:
45423
diff
changeset
|
120 |
fun decode_typ tm = |
fca432074fb2
sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents:
45423
diff
changeset
|
121 |
let |
fca432074fb2
sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents:
45423
diff
changeset
|
122 |
fun err () = raise TERM ("decode_typ: bad encoding of type", [tm]); |
42242
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
123 |
|
49685
4341e4d86872
allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents:
49674
diff
changeset
|
124 |
fun typ ps sort tm = |
4341e4d86872
allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents:
49674
diff
changeset
|
125 |
(case tm of |
4341e4d86872
allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents:
49674
diff
changeset
|
126 |
Const ("_tfree", _) $ t => typ ps sort t |
4341e4d86872
allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents:
49674
diff
changeset
|
127 |
| Const ("_tvar", _) $ t => typ ps sort t |
4341e4d86872
allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents:
49674
diff
changeset
|
128 |
| Const ("_ofsort", _) $ t $ s => |
4341e4d86872
allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents:
49674
diff
changeset
|
129 |
(case decode_pos s of |
4341e4d86872
allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents:
49674
diff
changeset
|
130 |
SOME p => typ (p :: ps) sort t |
4341e4d86872
allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents:
49674
diff
changeset
|
131 |
| NONE => |
4341e4d86872
allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents:
49674
diff
changeset
|
132 |
if is_none sort then typ ps (SOME (decode_sort s)) t |
4341e4d86872
allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents:
49674
diff
changeset
|
133 |
else err ()) |
4341e4d86872
allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents:
49674
diff
changeset
|
134 |
| Const ("_dummy_ofsort", _) $ s => TFree ("'_dummy_", decode_sort s) |
4341e4d86872
allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents:
49674
diff
changeset
|
135 |
| Free (x, _) => TFree (x, ps @ the_default dummyS sort) |
4341e4d86872
allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents:
49674
diff
changeset
|
136 |
| Var (xi, _) => TVar (xi, ps @ the_default dummyS sort) |
4341e4d86872
allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents:
49674
diff
changeset
|
137 |
| _ => |
4341e4d86872
allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents:
49674
diff
changeset
|
138 |
if null ps andalso is_none sort then |
4341e4d86872
allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents:
49674
diff
changeset
|
139 |
let |
4341e4d86872
allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents:
49674
diff
changeset
|
140 |
val (head, args) = Term.strip_comb tm; |
4341e4d86872
allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents:
49674
diff
changeset
|
141 |
val a = |
4341e4d86872
allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents:
49674
diff
changeset
|
142 |
(case head of |
4341e4d86872
allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents:
49674
diff
changeset
|
143 |
Const (c, _) => (Lexicon.unmark_type c handle Fail _ => err ()) |
4341e4d86872
allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents:
49674
diff
changeset
|
144 |
| _ => err ()); |
4341e4d86872
allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents:
49674
diff
changeset
|
145 |
in Type (a, map (typ [] NONE) args) end |
4341e4d86872
allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents:
49674
diff
changeset
|
146 |
else err ()); |
4341e4d86872
allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents:
49674
diff
changeset
|
147 |
in typ [] NONE tm end; |
42242
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
148 |
|
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
149 |
|
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
150 |
(* parsetree_to_ast *) |
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
151 |
|
81233
0199acc01aa8
more parser markup, based on position constraints for logical mixfix syntax;
wenzelm
parents:
81232
diff
changeset
|
152 |
fun parsetree_literals (Parser.Markup (_, ts)) = maps parsetree_literals ts |
0199acc01aa8
more parser markup, based on position constraints for logical mixfix syntax;
wenzelm
parents:
81232
diff
changeset
|
153 |
| parsetree_literals (Parser.Node _) = [] |
0199acc01aa8
more parser markup, based on position constraints for logical mixfix syntax;
wenzelm
parents:
81232
diff
changeset
|
154 |
| parsetree_literals (Parser.Tip tok) = |
0199acc01aa8
more parser markup, based on position constraints for logical mixfix syntax;
wenzelm
parents:
81232
diff
changeset
|
155 |
if Lexicon.is_literal tok andalso |
0199acc01aa8
more parser markup, based on position constraints for logical mixfix syntax;
wenzelm
parents:
81232
diff
changeset
|
156 |
not (null (Lexicon.literal_markup (Lexicon.str_of_token tok))) |
81251
89ea66c2045b
clarified position constraints: omit redundant information, e.g. for implicit object-logic judgment;
wenzelm
parents:
81250
diff
changeset
|
157 |
then filter Position.is_reported [Lexicon.pos_of_token tok] else []; |
81233
0199acc01aa8
more parser markup, based on position constraints for logical mixfix syntax;
wenzelm
parents:
81232
diff
changeset
|
158 |
|
52188 | 159 |
fun parsetree_to_ast ctxt trf parsetree = |
42242
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
160 |
let |
55922
710bc66f432c
more markup for inner syntax class/type names (notably for completion);
wenzelm
parents:
55829
diff
changeset
|
161 |
val reports = Unsynchronized.ref ([]: Position.report_text list); |
81558
b57996a0688c
clarified term positions and markup: syntax = true means this is via concrete syntax;
wenzelm
parents:
81552
diff
changeset
|
162 |
fun report pos = Term_Position.store_reports reports [pos]; |
55959
c3b458435f4f
more decisive commitment to get_free vs. the_const;
wenzelm
parents:
55957
diff
changeset
|
163 |
val append_reports = Position.append_reports reports; |
42242
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
164 |
|
73163
624c2b98860a
suppress markup for literal tokens with block control symbols, for better PIDE/HTML output (see also d15fe10593ff);
wenzelm
parents:
71675
diff
changeset
|
165 |
fun report_pos tok = |
80952
a61ed25ba155
clarified persistent datatype: more direct literal_markup, which also serves as a flag;
wenzelm
parents:
80950
diff
changeset
|
166 |
if Lexicon.is_literal tok andalso null (Lexicon.literal_markup (Lexicon.str_of_token tok)) |
a61ed25ba155
clarified persistent datatype: more direct literal_markup, which also serves as a flag;
wenzelm
parents:
80950
diff
changeset
|
167 |
then Position.none else Lexicon.pos_of_token tok; |
73163
624c2b98860a
suppress markup for literal tokens with block control symbols, for better PIDE/HTML output (see also d15fe10593ff);
wenzelm
parents:
71675
diff
changeset
|
168 |
|
81552 | 169 |
val markup_cache = Symtab.apply_unsynchronized_cache (markup_entity ctxt); |
81173
6002cb6bfb0a
performance tuning: cache markup and extern operations;
wenzelm
parents:
81169
diff
changeset
|
170 |
|
81211 | 171 |
val ast_of_pos = Ast.Variable o Term_Position.encode; |
81558
b57996a0688c
clarified term positions and markup: syntax = true means this is via concrete syntax;
wenzelm
parents:
81552
diff
changeset
|
172 |
val ast_of_position = Ast.Variable o Term_Position.encode_no_syntax o single o report_pos; |
b57996a0688c
clarified term positions and markup: syntax = true means this is via concrete syntax;
wenzelm
parents:
81552
diff
changeset
|
173 |
|
b57996a0688c
clarified term positions and markup: syntax = true means this is via concrete syntax;
wenzelm
parents:
81552
diff
changeset
|
174 |
val syntax_ast_of_pos = Ast.Variable o Term_Position.encode_syntax; |
b57996a0688c
clarified term positions and markup: syntax = true means this is via concrete syntax;
wenzelm
parents:
81552
diff
changeset
|
175 |
val syntax_ast_of_position = syntax_ast_of_pos o single o report_pos; |
b57996a0688c
clarified term positions and markup: syntax = true means this is via concrete syntax;
wenzelm
parents:
81552
diff
changeset
|
176 |
fun syntax_ast_of_position' a = Ast.constrain (Ast.Constant a) o syntax_ast_of_position; |
81211 | 177 |
|
49674
dbadb4d03cbc
report sort assignment of visible type variables;
wenzelm
parents:
49662
diff
changeset
|
178 |
fun asts_of_token tok = |
dbadb4d03cbc
report sort assignment of visible type variables;
wenzelm
parents:
49662
diff
changeset
|
179 |
if Lexicon.valued_token tok |
dbadb4d03cbc
report sort assignment of visible type variables;
wenzelm
parents:
49662
diff
changeset
|
180 |
then [Ast.Variable (Lexicon.str_of_token tok)] |
dbadb4d03cbc
report sort assignment of visible type variables;
wenzelm
parents:
49662
diff
changeset
|
181 |
else []; |
dbadb4d03cbc
report sort assignment of visible type variables;
wenzelm
parents:
49662
diff
changeset
|
182 |
|
dbadb4d03cbc
report sort assignment of visible type variables;
wenzelm
parents:
49662
diff
changeset
|
183 |
fun asts_of_position c tok = |
52188 | 184 |
[Ast.Appl [Ast.Constant c, ast_of (Parser.Tip tok), ast_of_position tok]] |
49674
dbadb4d03cbc
report sort assignment of visible type variables;
wenzelm
parents:
49662
diff
changeset
|
185 |
|
81008
d0cd220d8e8b
more markup reports: notably "notation=..." within pretty blocks;
wenzelm
parents:
81003
diff
changeset
|
186 |
and asts_of (Parser.Markup ({markup, range = (pos, _), ...}, pts)) = |
81016
8e2114e6205b
clarified order of markup: more uniform input vs. output;
wenzelm
parents:
81009
diff
changeset
|
187 |
let |
8e2114e6205b
clarified order of markup: more uniform input vs. output;
wenzelm
parents:
81009
diff
changeset
|
188 |
val asts = maps asts_of pts; |
8e2114e6205b
clarified order of markup: more uniform input vs. output;
wenzelm
parents:
81009
diff
changeset
|
189 |
val _ = append_reports (map (pair pos) markup); |
8e2114e6205b
clarified order of markup: more uniform input vs. output;
wenzelm
parents:
81009
diff
changeset
|
190 |
in asts end |
81001
0c6a600c8939
more detailed parse tree: retain nonterminal context as well;
wenzelm
parents:
80999
diff
changeset
|
191 |
| asts_of (Parser.Node ({const = "", ...}, pts)) = maps asts_of pts |
0c6a600c8939
more detailed parse tree: retain nonterminal context as well;
wenzelm
parents:
80999
diff
changeset
|
192 |
| asts_of (Parser.Node ({const = "_class_name", ...}, [Parser.Tip tok])) = |
42242
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
193 |
let |
73163
624c2b98860a
suppress markup for literal tokens with block control symbols, for better PIDE/HTML output (see also d15fe10593ff);
wenzelm
parents:
71675
diff
changeset
|
194 |
val pos = report_pos tok; |
55922
710bc66f432c
more markup for inner syntax class/type names (notably for completion);
wenzelm
parents:
55829
diff
changeset
|
195 |
val (c, rs) = Proof_Context.check_class ctxt (Lexicon.str_of_token tok, pos); |
55959
c3b458435f4f
more decisive commitment to get_free vs. the_const;
wenzelm
parents:
55957
diff
changeset
|
196 |
val _ = append_reports rs; |
42282
4fa41e068ff0
report literal tokens according to parsetree head;
wenzelm
parents:
42281
diff
changeset
|
197 |
in [Ast.Constant (Lexicon.mark_class c)] end |
81001
0c6a600c8939
more detailed parse tree: retain nonterminal context as well;
wenzelm
parents:
80999
diff
changeset
|
198 |
| asts_of (Parser.Node ({const = "_type_name", ...}, [Parser.Tip tok])) = |
42242
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
199 |
let |
73163
624c2b98860a
suppress markup for literal tokens with block control symbols, for better PIDE/HTML output (see also d15fe10593ff);
wenzelm
parents:
71675
diff
changeset
|
200 |
val pos = report_pos tok; |
80997 | 201 |
val (c, rs) = |
56002 | 202 |
Proof_Context.check_type_name {proper = true, strict = false} ctxt |
80997 | 203 |
(Lexicon.str_of_token tok, pos) |
204 |
|>> dest_Type_name; |
|
55959
c3b458435f4f
more decisive commitment to get_free vs. the_const;
wenzelm
parents:
55957
diff
changeset
|
205 |
val _ = append_reports rs; |
42282
4fa41e068ff0
report literal tokens according to parsetree head;
wenzelm
parents:
42281
diff
changeset
|
206 |
in [Ast.Constant (Lexicon.mark_type c)] end |
81279 | 207 |
| asts_of (Parser.Node ({const = "_DDDOT", ...}, [Parser.Tip tok])) = |
81558
b57996a0688c
clarified term positions and markup: syntax = true means this is via concrete syntax;
wenzelm
parents:
81552
diff
changeset
|
208 |
[Ast.constrain (Ast.Variable Auto_Bind.dddot_vname) (syntax_ast_of_position tok)] |
81001
0c6a600c8939
more detailed parse tree: retain nonterminal context as well;
wenzelm
parents:
80999
diff
changeset
|
209 |
| asts_of (Parser.Node ({const = "_position", ...}, [Parser.Tip tok])) = |
0c6a600c8939
more detailed parse tree: retain nonterminal context as well;
wenzelm
parents:
80999
diff
changeset
|
210 |
asts_of_position "_constrain" tok |
0c6a600c8939
more detailed parse tree: retain nonterminal context as well;
wenzelm
parents:
80999
diff
changeset
|
211 |
| asts_of (Parser.Node ({const = "_position_sort", ...}, [Parser.Tip tok])) = |
0c6a600c8939
more detailed parse tree: retain nonterminal context as well;
wenzelm
parents:
80999
diff
changeset
|
212 |
asts_of_position "_ofsort" tok |
0c6a600c8939
more detailed parse tree: retain nonterminal context as well;
wenzelm
parents:
80999
diff
changeset
|
213 |
| asts_of (Parser.Node ({const = a as "\<^const>Pure.dummy_pattern", ...}, [Parser.Tip tok])) = |
81558
b57996a0688c
clarified term positions and markup: syntax = true means this is via concrete syntax;
wenzelm
parents:
81552
diff
changeset
|
214 |
[syntax_ast_of_position' a tok] |
81001
0c6a600c8939
more detailed parse tree: retain nonterminal context as well;
wenzelm
parents:
80999
diff
changeset
|
215 |
| asts_of (Parser.Node ({const = a as "_idtdummy", ...}, [Parser.Tip tok])) = |
81558
b57996a0688c
clarified term positions and markup: syntax = true means this is via concrete syntax;
wenzelm
parents:
81552
diff
changeset
|
216 |
[syntax_ast_of_position' a tok] |
81001
0c6a600c8939
more detailed parse tree: retain nonterminal context as well;
wenzelm
parents:
80999
diff
changeset
|
217 |
| asts_of (Parser.Node ({const = "_idtypdummy", ...}, pts as [Parser.Tip tok, _, _])) = |
81211 | 218 |
let val args = maps asts_of pts |
81558
b57996a0688c
clarified term positions and markup: syntax = true means this is via concrete syntax;
wenzelm
parents:
81552
diff
changeset
|
219 |
in [Ast.Appl (Ast.Constant "_constrain" :: syntax_ast_of_position' "_idtdummy" tok :: args)] end |
81001
0c6a600c8939
more detailed parse tree: retain nonterminal context as well;
wenzelm
parents:
80999
diff
changeset
|
220 |
| asts_of (Parser.Node ({const = a, ...}, pts)) = |
81233
0199acc01aa8
more parser markup, based on position constraints for logical mixfix syntax;
wenzelm
parents:
81232
diff
changeset
|
221 |
let |
81558
b57996a0688c
clarified term positions and markup: syntax = true means this is via concrete syntax;
wenzelm
parents:
81552
diff
changeset
|
222 |
val ps = map Term_Position.syntax (maps parsetree_literals pts); |
81233
0199acc01aa8
more parser markup, based on position constraints for logical mixfix syntax;
wenzelm
parents:
81232
diff
changeset
|
223 |
val args = maps asts_of pts; |
0199acc01aa8
more parser markup, based on position constraints for logical mixfix syntax;
wenzelm
parents:
81232
diff
changeset
|
224 |
fun head () = |
81251
89ea66c2045b
clarified position constraints: omit redundant information, e.g. for implicit object-logic judgment;
wenzelm
parents:
81250
diff
changeset
|
225 |
if not (null ps) andalso (Lexicon.is_fixed a orelse Lexicon.is_const a) |
81233
0199acc01aa8
more parser markup, based on position constraints for logical mixfix syntax;
wenzelm
parents:
81232
diff
changeset
|
226 |
then Ast.constrain (Ast.Constant a) (ast_of_pos ps) |
0199acc01aa8
more parser markup, based on position constraints for logical mixfix syntax;
wenzelm
parents:
81232
diff
changeset
|
227 |
else Ast.Constant a; |
0199acc01aa8
more parser markup, based on position constraints for logical mixfix syntax;
wenzelm
parents:
81232
diff
changeset
|
228 |
val _ = List.app (fn pos => report pos markup_cache a) ps; |
0199acc01aa8
more parser markup, based on position constraints for logical mixfix syntax;
wenzelm
parents:
81232
diff
changeset
|
229 |
in |
0199acc01aa8
more parser markup, based on position constraints for logical mixfix syntax;
wenzelm
parents:
81232
diff
changeset
|
230 |
[case trf a of |
0199acc01aa8
more parser markup, based on position constraints for logical mixfix syntax;
wenzelm
parents:
81232
diff
changeset
|
231 |
SOME f => f ctxt args |
0199acc01aa8
more parser markup, based on position constraints for logical mixfix syntax;
wenzelm
parents:
81232
diff
changeset
|
232 |
| NONE => Ast.mk_appl (head ()) args] |
0199acc01aa8
more parser markup, based on position constraints for logical mixfix syntax;
wenzelm
parents:
81232
diff
changeset
|
233 |
end |
49674
dbadb4d03cbc
report sort assignment of visible type variables;
wenzelm
parents:
49662
diff
changeset
|
234 |
| asts_of (Parser.Tip tok) = asts_of_token tok |
42282
4fa41e068ff0
report literal tokens according to parsetree head;
wenzelm
parents:
42281
diff
changeset
|
235 |
|
4fa41e068ff0
report literal tokens according to parsetree head;
wenzelm
parents:
42281
diff
changeset
|
236 |
and ast_of pt = |
4fa41e068ff0
report literal tokens according to parsetree head;
wenzelm
parents:
42281
diff
changeset
|
237 |
(case asts_of pt of |
4fa41e068ff0
report literal tokens according to parsetree head;
wenzelm
parents:
42281
diff
changeset
|
238 |
[ast] => ast |
4fa41e068ff0
report literal tokens according to parsetree head;
wenzelm
parents:
42281
diff
changeset
|
239 |
| asts => raise Ast.AST ("parsetree_to_ast: malformed parsetree", asts)); |
42242
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
240 |
|
78705
fde0b195cb7d
clarified signature: avoid association with potentially dangerous Exn.capture;
wenzelm
parents:
74571
diff
changeset
|
241 |
val ast = Exn.result ast_of parsetree; |
42242
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
242 |
in (! reports, ast) end; |
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
243 |
|
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
244 |
|
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
245 |
(* ast_to_term *) |
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
246 |
|
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
247 |
fun ast_to_term ctxt trf = |
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
248 |
let |
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
249 |
fun trans a args = |
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
250 |
(case trf a of |
42476 | 251 |
NONE => Term.list_comb (Syntax.const a, args) |
42242
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
252 |
| SOME f => f ctxt args); |
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
253 |
|
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
254 |
fun term_of (Ast.Constant a) = trans a [] |
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
255 |
| term_of (Ast.Variable x) = Lexicon.read_var x |
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
256 |
| term_of (Ast.Appl (Ast.Constant a :: (asts as _ :: _))) = |
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
257 |
trans a (map term_of asts) |
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
258 |
| term_of (Ast.Appl (ast :: (asts as _ :: _))) = |
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
259 |
Term.list_comb (term_of ast, map term_of asts) |
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
260 |
| term_of (ast as Ast.Appl _) = raise Ast.AST ("ast_to_term: malformed ast", [ast]); |
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
261 |
in term_of end; |
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
262 |
|
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
263 |
|
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
264 |
(* decode_term -- transform parse tree into raw term *) |
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
265 |
|
55959
c3b458435f4f
more decisive commitment to get_free vs. the_const;
wenzelm
parents:
55957
diff
changeset
|
266 |
fun decode_const ctxt (c, ps) = |
80997 | 267 |
Proof_Context.check_const {proper = true, strict = false} ctxt (c, ps) |
268 |
|>> dest_Const_name; |
|
55950
3bb5c7179234
clarified treatment of consts -- prefer value-oriented reports;
wenzelm
parents:
55949
diff
changeset
|
269 |
|
81173
6002cb6bfb0a
performance tuning: cache markup and extern operations;
wenzelm
parents:
81169
diff
changeset
|
270 |
fun decode_term ctxt = |
6002cb6bfb0a
performance tuning: cache markup and extern operations;
wenzelm
parents:
81169
diff
changeset
|
271 |
let |
81540 | 272 |
val markup_free_cache = Symtab.apply_unsynchronized_cache (markup_free ctxt); |
273 |
val markup_const_cache = Symtab.apply_unsynchronized_cache (markup_const ctxt); |
|
81173
6002cb6bfb0a
performance tuning: cache markup and extern operations;
wenzelm
parents:
81169
diff
changeset
|
274 |
|
6002cb6bfb0a
performance tuning: cache markup and extern operations;
wenzelm
parents:
81169
diff
changeset
|
275 |
fun decode (result as (_: Position.report_text list, Exn.Exn _)) = result |
6002cb6bfb0a
performance tuning: cache markup and extern operations;
wenzelm
parents:
81169
diff
changeset
|
276 |
| decode (reports0, Exn.Res tm) = |
6002cb6bfb0a
performance tuning: cache markup and extern operations;
wenzelm
parents:
81169
diff
changeset
|
277 |
let |
6002cb6bfb0a
performance tuning: cache markup and extern operations;
wenzelm
parents:
81169
diff
changeset
|
278 |
val reports = Unsynchronized.ref reports0; |
81558
b57996a0688c
clarified term positions and markup: syntax = true means this is via concrete syntax;
wenzelm
parents:
81552
diff
changeset
|
279 |
fun report ps = Term_Position.store_reports reports ps; |
81173
6002cb6bfb0a
performance tuning: cache markup and extern operations;
wenzelm
parents:
81169
diff
changeset
|
280 |
val append_reports = Position.append_reports reports; |
55959
c3b458435f4f
more decisive commitment to get_free vs. the_const;
wenzelm
parents:
55957
diff
changeset
|
281 |
|
81173
6002cb6bfb0a
performance tuning: cache markup and extern operations;
wenzelm
parents:
81169
diff
changeset
|
282 |
fun decode ps qs bs (Const ("_constrain", _) $ t $ typ) = |
6002cb6bfb0a
performance tuning: cache markup and extern operations;
wenzelm
parents:
81169
diff
changeset
|
283 |
(case Term_Position.decode_position typ of |
81218 | 284 |
SOME (ps', T) => Type.constraint T (decode (ps' @ ps) qs bs t) |
81173
6002cb6bfb0a
performance tuning: cache markup and extern operations;
wenzelm
parents:
81169
diff
changeset
|
285 |
| NONE => Type.constraint (decode_typ typ) (decode ps qs bs t)) |
6002cb6bfb0a
performance tuning: cache markup and extern operations;
wenzelm
parents:
81169
diff
changeset
|
286 |
| decode ps qs bs (Const ("_constrainAbs", _) $ t $ typ) = |
6002cb6bfb0a
performance tuning: cache markup and extern operations;
wenzelm
parents:
81169
diff
changeset
|
287 |
(case Term_Position.decode_position typ of |
81218 | 288 |
SOME (qs', T) => Type.constraint (T --> dummyT) (decode ps (qs' @ qs) bs t) |
81173
6002cb6bfb0a
performance tuning: cache markup and extern operations;
wenzelm
parents:
81169
diff
changeset
|
289 |
| NONE => Type.constraint (decode_typ typ --> dummyT) (decode ps qs bs t)) |
6002cb6bfb0a
performance tuning: cache markup and extern operations;
wenzelm
parents:
81169
diff
changeset
|
290 |
| decode _ qs bs (Abs (x, T, t)) = |
42242
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
291 |
let |
81173
6002cb6bfb0a
performance tuning: cache markup and extern operations;
wenzelm
parents:
81169
diff
changeset
|
292 |
val id = serial (); |
6002cb6bfb0a
performance tuning: cache markup and extern operations;
wenzelm
parents:
81169
diff
changeset
|
293 |
val _ = report qs (markup_bound {def = true} qs) (x, id); |
6002cb6bfb0a
performance tuning: cache markup and extern operations;
wenzelm
parents:
81169
diff
changeset
|
294 |
in Abs (x, T, decode [] [] ((qs, (x, id)) :: bs) t) end |
6002cb6bfb0a
performance tuning: cache markup and extern operations;
wenzelm
parents:
81169
diff
changeset
|
295 |
| decode _ _ bs (t $ u) = decode [] [] bs t $ decode [] [] bs u |
6002cb6bfb0a
performance tuning: cache markup and extern operations;
wenzelm
parents:
81169
diff
changeset
|
296 |
| decode ps _ _ (Const (a, T)) = |
6002cb6bfb0a
performance tuning: cache markup and extern operations;
wenzelm
parents:
81169
diff
changeset
|
297 |
(case try Lexicon.unmark_fixed a of |
81233
0199acc01aa8
more parser markup, based on position constraints for logical mixfix syntax;
wenzelm
parents:
81232
diff
changeset
|
298 |
SOME x => Free (x, T) |
81173
6002cb6bfb0a
performance tuning: cache markup and extern operations;
wenzelm
parents:
81169
diff
changeset
|
299 |
| NONE => |
81233
0199acc01aa8
more parser markup, based on position constraints for logical mixfix syntax;
wenzelm
parents:
81232
diff
changeset
|
300 |
(case try Lexicon.unmark_const a of |
0199acc01aa8
more parser markup, based on position constraints for logical mixfix syntax;
wenzelm
parents:
81232
diff
changeset
|
301 |
SOME c => Const (c, T) |
0199acc01aa8
more parser markup, based on position constraints for logical mixfix syntax;
wenzelm
parents:
81232
diff
changeset
|
302 |
| NONE => |
0199acc01aa8
more parser markup, based on position constraints for logical mixfix syntax;
wenzelm
parents:
81232
diff
changeset
|
303 |
let val c = #1 (decode_const ctxt (a, [])) |
0199acc01aa8
more parser markup, based on position constraints for logical mixfix syntax;
wenzelm
parents:
81232
diff
changeset
|
304 |
in report ps markup_const_cache c; Const (c, T) end)) |
81173
6002cb6bfb0a
performance tuning: cache markup and extern operations;
wenzelm
parents:
81169
diff
changeset
|
305 |
| decode ps _ _ (Free (a, T)) = |
81558
b57996a0688c
clarified term positions and markup: syntax = true means this is via concrete syntax;
wenzelm
parents:
81552
diff
changeset
|
306 |
((Name.reject_internal (a, map #pos ps) handle ERROR msg => |
b57996a0688c
clarified term positions and markup: syntax = true means this is via concrete syntax;
wenzelm
parents:
81552
diff
changeset
|
307 |
error (msg ^ Proof_Context.consts_completion_message ctxt (a, map #pos ps))); |
81537
d230683a35fc
more direct Proof_Context.lookup_free -- bypass redundancy of Proof_Context.check_const;
wenzelm
parents:
81279
diff
changeset
|
308 |
(case Proof_Context.lookup_free ctxt a of |
81173
6002cb6bfb0a
performance tuning: cache markup and extern operations;
wenzelm
parents:
81169
diff
changeset
|
309 |
SOME x => (report ps markup_free_cache x; Free (x, T)) |
6002cb6bfb0a
performance tuning: cache markup and extern operations;
wenzelm
parents:
81169
diff
changeset
|
310 |
| NONE => |
6002cb6bfb0a
performance tuning: cache markup and extern operations;
wenzelm
parents:
81169
diff
changeset
|
311 |
let |
81558
b57996a0688c
clarified term positions and markup: syntax = true means this is via concrete syntax;
wenzelm
parents:
81552
diff
changeset
|
312 |
val (c, rs) = decode_const ctxt (a, map #pos ps); |
81173
6002cb6bfb0a
performance tuning: cache markup and extern operations;
wenzelm
parents:
81169
diff
changeset
|
313 |
val _ = append_reports rs; |
6002cb6bfb0a
performance tuning: cache markup and extern operations;
wenzelm
parents:
81169
diff
changeset
|
314 |
in Const (c, T) end)) |
6002cb6bfb0a
performance tuning: cache markup and extern operations;
wenzelm
parents:
81169
diff
changeset
|
315 |
| decode ps _ _ (Var (xi, T)) = (report ps markup_var xi; Var (xi, T)) |
6002cb6bfb0a
performance tuning: cache markup and extern operations;
wenzelm
parents:
81169
diff
changeset
|
316 |
| decode ps _ bs (t as Bound i) = |
6002cb6bfb0a
performance tuning: cache markup and extern operations;
wenzelm
parents:
81169
diff
changeset
|
317 |
(case try (nth bs) i of |
6002cb6bfb0a
performance tuning: cache markup and extern operations;
wenzelm
parents:
81169
diff
changeset
|
318 |
SOME (qs, (x, id)) => (report ps (markup_bound {def = false} qs) (x, id); t) |
6002cb6bfb0a
performance tuning: cache markup and extern operations;
wenzelm
parents:
81169
diff
changeset
|
319 |
| NONE => t); |
42242
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
320 |
|
81173
6002cb6bfb0a
performance tuning: cache markup and extern operations;
wenzelm
parents:
81169
diff
changeset
|
321 |
val tm' = Exn.result (fn () => decode [] [] [] tm) (); |
6002cb6bfb0a
performance tuning: cache markup and extern operations;
wenzelm
parents:
81169
diff
changeset
|
322 |
in (! reports, tm') end; |
6002cb6bfb0a
performance tuning: cache markup and extern operations;
wenzelm
parents:
81169
diff
changeset
|
323 |
|
6002cb6bfb0a
performance tuning: cache markup and extern operations;
wenzelm
parents:
81169
diff
changeset
|
324 |
in decode end; |
42242
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
325 |
|
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
326 |
|
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
327 |
|
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
328 |
(** parse **) |
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
329 |
|
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
330 |
(* results *) |
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
331 |
|
43761
e72ba84ae58f
tuned signature -- corresponding to Scala version;
wenzelm
parents:
43731
diff
changeset
|
332 |
fun proper_results results = map_filter (fn (y, Exn.Res x) => SOME (y, x) | _ => NONE) results; |
42242
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
333 |
fun failed_results results = map_filter (fn (y, Exn.Exn e) => SOME (y, e) | _ => NONE) results; |
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
334 |
|
46512
4f9f61f9b535
simplified configuration options for syntax ambiguity;
wenzelm
parents:
46506
diff
changeset
|
335 |
fun report_result ctxt pos ambig_msgs results = |
42242
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
336 |
(case (proper_results results, failed_results results) of |
62505 | 337 |
([], (reports, exn) :: _) => (Context_Position.reports_text ctxt reports; Exn.reraise exn) |
55922
710bc66f432c
more markup for inner syntax class/type names (notably for completion);
wenzelm
parents:
55829
diff
changeset
|
338 |
| ([(reports, x)], _) => (Context_Position.reports_text ctxt reports; x) |
46512
4f9f61f9b535
simplified configuration options for syntax ambiguity;
wenzelm
parents:
46506
diff
changeset
|
339 |
| _ => |
4f9f61f9b535
simplified configuration options for syntax ambiguity;
wenzelm
parents:
46506
diff
changeset
|
340 |
if null ambig_msgs then |
48992 | 341 |
error ("Parse error: ambiguous syntax" ^ Position.here pos) |
46512
4f9f61f9b535
simplified configuration options for syntax ambiguity;
wenzelm
parents:
46506
diff
changeset
|
342 |
else error (cat_lines ambig_msgs)); |
42242
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
343 |
|
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
344 |
|
42249 | 345 |
(* parse raw asts *) |
42242
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
346 |
|
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
347 |
fun parse_asts ctxt raw root (syms, pos) = |
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
348 |
let |
80074
951c371c1cd9
clarified names: discontinue odd convention from 3 decades ago;
wenzelm
parents:
78705
diff
changeset
|
349 |
val syn = Proof_Context.syntax_of ctxt; |
42253 | 350 |
val ast_tr = Syntax.parse_ast_translation syn; |
42242
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
351 |
|
42251 | 352 |
val toks = Syntax.tokenize syn raw syms; |
73163
624c2b98860a
suppress markup for literal tokens with block control symbols, for better PIDE/HTML output (see also d15fe10593ff);
wenzelm
parents:
71675
diff
changeset
|
353 |
val _ = Context_Position.reports ctxt (maps Lexicon.reports_of_token toks); |
42242
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
354 |
|
45641 | 355 |
val pts = Syntax.parse syn root (filter Lexicon.is_proper toks) |
42242
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
356 |
handle ERROR msg => |
55957
cffb46aea3d1
more compact Markup.markup_report: message body may consist of multiple elements;
wenzelm
parents:
55956
diff
changeset
|
357 |
error (msg ^ Markup.markup_report (implode (map (Lexicon.reported_token_range ctxt) toks))); |
42242
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
358 |
val len = length pts; |
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
359 |
|
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
360 |
val limit = Config.get ctxt Syntax.ambiguity_limit; |
46512
4f9f61f9b535
simplified configuration options for syntax ambiguity;
wenzelm
parents:
46506
diff
changeset
|
361 |
val ambig_msgs = |
4f9f61f9b535
simplified configuration options for syntax ambiguity;
wenzelm
parents:
46506
diff
changeset
|
362 |
if len <= 1 then [] |
46506
c7faa011bfa7
simplified configuration options for syntax ambiguity;
wenzelm
parents:
46126
diff
changeset
|
363 |
else |
46512
4f9f61f9b535
simplified configuration options for syntax ambiguity;
wenzelm
parents:
46506
diff
changeset
|
364 |
[cat_lines |
62800 | 365 |
(("Ambiguous input" ^ Position.here (Position.no_range_position pos) ^ |
55979
06cb126f30ba
tuned message: reveal ambiguity where Syntax_Phases.decode_term fails and thus reduces proper_results beforehand, e.g. term "f(x := y)" in ~~/src/HOL/Hoare/Hoare_Logic.thy;
wenzelm
parents:
55976
diff
changeset
|
366 |
" produces " ^ string_of_int len ^ " parse trees" ^ |
46512
4f9f61f9b535
simplified configuration options for syntax ambiguity;
wenzelm
parents:
46506
diff
changeset
|
367 |
(if len <= limit then "" else " (" ^ string_of_int limit ^ " displayed)") ^ ":") :: |
80991 | 368 |
map (Pretty.string_of o Pretty.item o Parser.pretty_tree) (take limit pts))]; |
46512
4f9f61f9b535
simplified configuration options for syntax ambiguity;
wenzelm
parents:
46506
diff
changeset
|
369 |
|
52188 | 370 |
in (ambig_msgs, map (parsetree_to_ast ctxt ast_tr) pts) end; |
42242
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
371 |
|
46506
c7faa011bfa7
simplified configuration options for syntax ambiguity;
wenzelm
parents:
46126
diff
changeset
|
372 |
fun parse_tree ctxt root input = |
42242
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
373 |
let |
80074
951c371c1cd9
clarified names: discontinue odd convention from 3 decades ago;
wenzelm
parents:
78705
diff
changeset
|
374 |
val syn = Proof_Context.syntax_of ctxt; |
42253 | 375 |
val tr = Syntax.parse_translation syn; |
81208
893b056542e7
more accurate treatment of constraints: restrict permissive mode to output;
wenzelm
parents:
81201
diff
changeset
|
376 |
val normalize = Ast.normalize ctxt {permissive_constraints = false} (Syntax.parse_rules syn); |
46512
4f9f61f9b535
simplified configuration options for syntax ambiguity;
wenzelm
parents:
46506
diff
changeset
|
377 |
val (ambig_msgs, asts) = parse_asts ctxt false root input; |
81208
893b056542e7
more accurate treatment of constraints: restrict permissive mode to output;
wenzelm
parents:
81201
diff
changeset
|
378 |
val results = (map o apsnd o Exn.maps_res) (normalize #> Exn.result (ast_to_term ctxt tr)) asts; |
46512
4f9f61f9b535
simplified configuration options for syntax ambiguity;
wenzelm
parents:
46506
diff
changeset
|
379 |
in (ambig_msgs, results) end; |
42242
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
380 |
|
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
381 |
|
42249 | 382 |
(* parse logical entities *) |
42242
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
383 |
|
42241
dd8029f71e1c
separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff
changeset
|
384 |
fun parse_failed ctxt pos msg kind = |
dd8029f71e1c
separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff
changeset
|
385 |
cat_error msg ("Failed to parse " ^ kind ^ |
64677
8dc24130e8fe
more uniform treatment of "bad" like other messages (with serial number);
wenzelm
parents:
63395
diff
changeset
|
386 |
Markup.markup_report (Context_Position.reported_text ctxt pos (Markup.bad ()) "")); |
42241
dd8029f71e1c
separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff
changeset
|
387 |
|
43731
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents:
43552
diff
changeset
|
388 |
fun parse_sort ctxt = |
59795 | 389 |
Syntax.parse_input ctxt Term_XML.Decode.sort Markup.language_sort |
43731
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents:
43552
diff
changeset
|
390 |
(fn (syms, pos) => |
46506
c7faa011bfa7
simplified configuration options for syntax ambiguity;
wenzelm
parents:
46126
diff
changeset
|
391 |
parse_tree ctxt "sort" (syms, pos) |
46512
4f9f61f9b535
simplified configuration options for syntax ambiguity;
wenzelm
parents:
46506
diff
changeset
|
392 |
|> uncurry (report_result ctxt pos) |
45427
fca432074fb2
sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents:
45423
diff
changeset
|
393 |
|> decode_sort |
43731
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents:
43552
diff
changeset
|
394 |
|> Type.minimize_sort (Proof_Context.tsig_of ctxt) |
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents:
43552
diff
changeset
|
395 |
handle ERROR msg => parse_failed ctxt pos msg "sort"); |
42241
dd8029f71e1c
separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff
changeset
|
396 |
|
43731
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents:
43552
diff
changeset
|
397 |
fun parse_typ ctxt = |
59795 | 398 |
Syntax.parse_input ctxt Term_XML.Decode.typ Markup.language_type |
43731
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents:
43552
diff
changeset
|
399 |
(fn (syms, pos) => |
46506
c7faa011bfa7
simplified configuration options for syntax ambiguity;
wenzelm
parents:
46126
diff
changeset
|
400 |
parse_tree ctxt "type" (syms, pos) |
46512
4f9f61f9b535
simplified configuration options for syntax ambiguity;
wenzelm
parents:
46506
diff
changeset
|
401 |
|> uncurry (report_result ctxt pos) |
45427
fca432074fb2
sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents:
45423
diff
changeset
|
402 |
|> decode_typ |
43731
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents:
43552
diff
changeset
|
403 |
handle ERROR msg => parse_failed ctxt pos msg "type"); |
42241
dd8029f71e1c
separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff
changeset
|
404 |
|
43731
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents:
43552
diff
changeset
|
405 |
fun parse_term is_prop ctxt = |
42241
dd8029f71e1c
separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff
changeset
|
406 |
let |
42281 | 407 |
val (markup, kind, root, constrain) = |
408 |
if is_prop |
|
55550 | 409 |
then (Markup.language_prop, "prop", "prop", Type.constraint propT) |
410 |
else (Markup.language_term, "term", Config.get ctxt Syntax.root, I); |
|
70784
799437173553
Term_XML.Encode/Decode.term uses Const "typargs";
wenzelm
parents:
69042
diff
changeset
|
411 |
val decode = constrain o Term_XML.Decode.term (Proof_Context.consts_of ctxt); |
42249 | 412 |
in |
59795 | 413 |
Syntax.parse_input ctxt decode markup |
43731
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents:
43552
diff
changeset
|
414 |
(fn (syms, pos) => |
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents:
43552
diff
changeset
|
415 |
let |
46512
4f9f61f9b535
simplified configuration options for syntax ambiguity;
wenzelm
parents:
46506
diff
changeset
|
416 |
val (ambig_msgs, results) = parse_tree ctxt root (syms, pos) ||> map (decode_term ctxt); |
46506
c7faa011bfa7
simplified configuration options for syntax ambiguity;
wenzelm
parents:
46126
diff
changeset
|
417 |
val parsed_len = length (proper_results results); |
42249 | 418 |
|
46512
4f9f61f9b535
simplified configuration options for syntax ambiguity;
wenzelm
parents:
46506
diff
changeset
|
419 |
val ambiguity_warning = Config.get ctxt Syntax.ambiguity_warning; |
43731
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents:
43552
diff
changeset
|
420 |
val limit = Config.get ctxt Syntax.ambiguity_limit; |
42249 | 421 |
|
43731
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents:
43552
diff
changeset
|
422 |
(*brute-force disambiguation via type-inference*) |
60490
9b28f446d9c5
more informative check: dummies are always allowed parse_term and should not lead to rejection here;
wenzelm
parents:
59795
diff
changeset
|
423 |
fun check t = |
9b28f446d9c5
more informative check: dummies are always allowed parse_term and should not lead to rejection here;
wenzelm
parents:
59795
diff
changeset
|
424 |
(Syntax.check_term (Proof_Context.allow_dummies ctxt) (constrain t); Exn.Res t) |
9b28f446d9c5
more informative check: dummies are always allowed parse_term and should not lead to rejection here;
wenzelm
parents:
59795
diff
changeset
|
425 |
handle exn as ERROR _ => Exn.Exn exn; |
42249 | 426 |
|
73686 | 427 |
fun par_map xs = Par_List.map' {name = "Syntax_Phases.parse_term", sequential = false} xs; |
43731
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents:
43552
diff
changeset
|
428 |
val results' = |
73686 | 429 |
if parsed_len > 1 then (grouped 10 par_map o apsnd o Exn.maps_res) check results |
43731
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents:
43552
diff
changeset
|
430 |
else results; |
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents:
43552
diff
changeset
|
431 |
val reports' = fst (hd results'); |
42249 | 432 |
|
43731
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents:
43552
diff
changeset
|
433 |
val errs = map snd (failed_results results'); |
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents:
43552
diff
changeset
|
434 |
val checked = map snd (proper_results results'); |
46506
c7faa011bfa7
simplified configuration options for syntax ambiguity;
wenzelm
parents:
46126
diff
changeset
|
435 |
val checked_len = length checked; |
42249 | 436 |
|
52517 | 437 |
val pretty_term = Syntax.pretty_term (Config.put Printer.show_brackets true ctxt); |
43731
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents:
43552
diff
changeset
|
438 |
in |
46506
c7faa011bfa7
simplified configuration options for syntax ambiguity;
wenzelm
parents:
46126
diff
changeset
|
439 |
if checked_len = 0 then |
46512
4f9f61f9b535
simplified configuration options for syntax ambiguity;
wenzelm
parents:
46506
diff
changeset
|
440 |
report_result ctxt pos [] |
4f9f61f9b535
simplified configuration options for syntax ambiguity;
wenzelm
parents:
46506
diff
changeset
|
441 |
[(reports', Exn.Exn (Exn.EXCEPTIONS (map ERROR ambig_msgs @ errs)))] |
46506
c7faa011bfa7
simplified configuration options for syntax ambiguity;
wenzelm
parents:
46126
diff
changeset
|
442 |
else if checked_len = 1 then |
56294
85911b8a6868
prefer Context_Position where a context is available;
wenzelm
parents:
56243
diff
changeset
|
443 |
(if not (null ambig_msgs) andalso ambiguity_warning andalso |
85911b8a6868
prefer Context_Position where a context is available;
wenzelm
parents:
56243
diff
changeset
|
444 |
Context_Position.is_visible ctxt then |
85911b8a6868
prefer Context_Position where a context is available;
wenzelm
parents:
56243
diff
changeset
|
445 |
warning |
46512
4f9f61f9b535
simplified configuration options for syntax ambiguity;
wenzelm
parents:
46506
diff
changeset
|
446 |
(cat_lines (ambig_msgs @ |
55979
06cb126f30ba
tuned message: reveal ambiguity where Syntax_Phases.decode_term fails and thus reduces proper_results beforehand, e.g. term "f(x := y)" in ~~/src/HOL/Hoare/Hoare_Logic.thy;
wenzelm
parents:
55976
diff
changeset
|
447 |
["Fortunately, only one parse tree is well-formed and type-correct,\n\ |
06cb126f30ba
tuned message: reveal ambiguity where Syntax_Phases.decode_term fails and thus reduces proper_results beforehand, e.g. term "f(x := y)" in ~~/src/HOL/Hoare/Hoare_Logic.thy;
wenzelm
parents:
55976
diff
changeset
|
448 |
\but you may still want to disambiguate your grammar or your input."])) |
46512
4f9f61f9b535
simplified configuration options for syntax ambiguity;
wenzelm
parents:
46506
diff
changeset
|
449 |
else (); report_result ctxt pos [] results') |
43731
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents:
43552
diff
changeset
|
450 |
else |
46512
4f9f61f9b535
simplified configuration options for syntax ambiguity;
wenzelm
parents:
46506
diff
changeset
|
451 |
report_result ctxt pos [] |
4f9f61f9b535
simplified configuration options for syntax ambiguity;
wenzelm
parents:
46506
diff
changeset
|
452 |
[(reports', Exn.Exn (ERROR (cat_lines (ambig_msgs @ |
4f9f61f9b535
simplified configuration options for syntax ambiguity;
wenzelm
parents:
46506
diff
changeset
|
453 |
(("Ambiguous input\n" ^ string_of_int checked_len ^ " terms are type correct" ^ |
46506
c7faa011bfa7
simplified configuration options for syntax ambiguity;
wenzelm
parents:
46126
diff
changeset
|
454 |
(if checked_len <= limit then "" |
c7faa011bfa7
simplified configuration options for syntax ambiguity;
wenzelm
parents:
46126
diff
changeset
|
455 |
else " (" ^ string_of_int limit ^ " displayed)") ^ ":") :: |
52517 | 456 |
map (Pretty.string_of o Pretty.item o single o pretty_term) |
457 |
(take limit checked))))))] |
|
43731
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents:
43552
diff
changeset
|
458 |
end handle ERROR msg => parse_failed ctxt pos msg kind) |
42249 | 459 |
end; |
42241
dd8029f71e1c
separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff
changeset
|
460 |
|
dd8029f71e1c
separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff
changeset
|
461 |
|
42242
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
462 |
(* parse_ast_pattern *) |
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
463 |
|
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
464 |
fun parse_ast_pattern ctxt (root, str) = |
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
465 |
let |
80074
951c371c1cd9
clarified names: discontinue odd convention from 3 decades ago;
wenzelm
parents:
78705
diff
changeset
|
466 |
val syn = Proof_Context.syntax_of ctxt; |
42242
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
467 |
|
55922
710bc66f432c
more markup for inner syntax class/type names (notably for completion);
wenzelm
parents:
55829
diff
changeset
|
468 |
val reports = Unsynchronized.ref ([]: Position.report_text list); |
81558
b57996a0688c
clarified term positions and markup: syntax = true means this is via concrete syntax;
wenzelm
parents:
81552
diff
changeset
|
469 |
fun report ps = Term_Position.store_reports reports ps; |
52188 | 470 |
|
81552 | 471 |
val markup_cache = Symtab.apply_unsynchronized_cache (markup_entity ctxt); |
81173
6002cb6bfb0a
performance tuning: cache markup and extern operations;
wenzelm
parents:
81169
diff
changeset
|
472 |
|
6002cb6bfb0a
performance tuning: cache markup and extern operations;
wenzelm
parents:
81169
diff
changeset
|
473 |
fun decode_const ps c = (report ps markup_cache c; Ast.Constant c); |
52188 | 474 |
fun decode_var ps x = (report ps (fn () => [Markup.name x Markup.free]) (); Ast.Variable x); |
475 |
fun decode_appl ps asts = Ast.Appl (map (decode ps) asts) |
|
476 |
and decode ps (Ast.Constant c) = decode_const ps c |
|
477 |
| decode ps (Ast.Variable x) = |
|
80739 | 478 |
if Syntax.is_const syn x orelse Long_Name.is_qualified x |
52188 | 479 |
then decode_const ps x |
480 |
else decode_var ps x |
|
481 |
| decode ps (Ast.Appl (asts as (Ast.Constant c :: ast :: Ast.Variable x :: args))) = |
|
482 |
if member (op =) Term_Position.markers c then |
|
483 |
(case Term_Position.decode x of |
|
81218 | 484 |
[] => decode_appl ps asts |
485 |
| ps' => Ast.mk_appl (decode (ps' @ ps) ast) (map (decode ps) args)) |
|
52188 | 486 |
else decode_appl ps asts |
487 |
| decode ps (Ast.Appl asts) = decode_appl ps asts; |
|
42242
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
488 |
|
59795 | 489 |
val source = Syntax.read_input str; |
59064 | 490 |
val pos = Input.pos_of source; |
491 |
val syms = Input.source_explode source; |
|
52188 | 492 |
val ast = |
58978
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
wenzelm
parents:
58047
diff
changeset
|
493 |
parse_asts ctxt true root (syms, pos) |
e42da880c61e
more position information, e.g. relevant for errors in generated ML source;
wenzelm
parents:
58047
diff
changeset
|
494 |
|> uncurry (report_result ctxt pos) |
52188 | 495 |
|> decode []; |
55922
710bc66f432c
more markup for inner syntax class/type names (notably for completion);
wenzelm
parents:
55829
diff
changeset
|
496 |
val _ = Context_Position.reports_text ctxt (! reports); |
52188 | 497 |
in ast end; |
42242
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
498 |
|
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
499 |
|
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
500 |
|
42245 | 501 |
(** encode parse trees **) |
502 |
||
503 |
(* term_of_sort *) |
|
504 |
||
505 |
fun term_of_sort S = |
|
506 |
let |
|
42476 | 507 |
val class = Syntax.const o Lexicon.mark_class; |
42245 | 508 |
|
509 |
fun classes [c] = class c |
|
42476 | 510 |
| classes (c :: cs) = Syntax.const "_classes" $ class c $ classes cs; |
42245 | 511 |
in |
67718 | 512 |
if S = dummyS then Syntax.const "_dummy_sort" |
513 |
else |
|
514 |
(case S of |
|
515 |
[] => Syntax.const "_topsort" |
|
516 |
| [c] => class c |
|
517 |
| cs => Syntax.const "_sort" $ classes cs) |
|
42245 | 518 |
end; |
519 |
||
520 |
||
521 |
(* term_of_typ *) |
|
522 |
||
42247
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
wenzelm
parents:
42245
diff
changeset
|
523 |
fun term_of_typ ctxt ty = |
42245 | 524 |
let |
52210
0226035df99d
more explicit Printer.type_emphasis, depending on show_type_emphasis;
wenzelm
parents:
52188
diff
changeset
|
525 |
val show_sorts = Config.get ctxt show_sorts orelse Config.get ctxt show_markup; |
42247
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
wenzelm
parents:
42245
diff
changeset
|
526 |
|
49686
678aa92e921c
printed "_ofsort" is subject to show_markup as well;
wenzelm
parents:
49685
diff
changeset
|
527 |
fun ofsort t raw_S = |
52210
0226035df99d
more explicit Printer.type_emphasis, depending on show_type_emphasis;
wenzelm
parents:
52188
diff
changeset
|
528 |
if show_sorts then |
49686
678aa92e921c
printed "_ofsort" is subject to show_markup as well;
wenzelm
parents:
49685
diff
changeset
|
529 |
let val S = #2 (Term_Position.decode_positionS raw_S) |
678aa92e921c
printed "_ofsort" is subject to show_markup as well;
wenzelm
parents:
49685
diff
changeset
|
530 |
in if S = dummyS then t else Syntax.const "_ofsort" $ t $ term_of_sort S end |
42245 | 531 |
else t; |
532 |
||
533 |
fun term_of (Type (a, Ts)) = |
|
42476 | 534 |
Term.list_comb (Syntax.const (Lexicon.mark_type a), map term_of Ts) |
42245 | 535 |
| term_of (TFree (x, S)) = |
81232 | 536 |
if Term_Position.detect x then Syntax.free x |
49686
678aa92e921c
printed "_ofsort" is subject to show_markup as well;
wenzelm
parents:
49685
diff
changeset
|
537 |
else ofsort (Syntax.const "_tfree" $ Syntax.free x) S |
678aa92e921c
printed "_ofsort" is subject to show_markup as well;
wenzelm
parents:
49685
diff
changeset
|
538 |
| term_of (TVar (xi, S)) = ofsort (Syntax.const "_tvar" $ Syntax.var xi) S; |
42245 | 539 |
in term_of ty end; |
540 |
||
541 |
||
542 |
(* simple_ast_of *) |
|
543 |
||
544 |
fun simple_ast_of ctxt = |
|
545 |
let |
|
546 |
val tune_var = if Config.get ctxt show_question_marks then I else unprefix "?"; |
|
547 |
fun ast_of (Const (c, _)) = Ast.Constant c |
|
548 |
| ast_of (Free (x, _)) = Ast.Variable x |
|
549 |
| ast_of (Var (xi, _)) = Ast.Variable (tune_var (Term.string_of_vname xi)) |
|
550 |
| ast_of (t as _ $ _) = |
|
551 |
let val (f, args) = strip_comb t |
|
552 |
in Ast.mk_appl (ast_of f) (map ast_of args) end |
|
42408 | 553 |
| ast_of (Bound i) = Ast.Appl [Ast.Constant "_loose", Ast.Variable ("B." ^ string_of_int i)] |
42245 | 554 |
| ast_of (Abs _) = raise Fail "simple_ast_of: Abs"; |
555 |
in ast_of end; |
|
556 |
||
557 |
||
558 |
(* sort_to_ast and typ_to_ast *) |
|
559 |
||
560 |
fun ast_of_termT ctxt trf tm = |
|
561 |
let |
|
562 |
val ctxt' = Config.put show_sorts false ctxt; |
|
563 |
fun ast_of (t as Const ("_tfree", _) $ Free _) = simple_ast_of ctxt t |
|
564 |
| ast_of (t as Const ("_tvar", _) $ Var _) = simple_ast_of ctxt t |
|
565 |
| ast_of (Const (a, _)) = trans a [] |
|
566 |
| ast_of (t as _ $ _) = |
|
567 |
(case strip_comb t of |
|
568 |
(Const (a, _), args) => trans a args |
|
569 |
| (f, args) => Ast.Appl (map ast_of (f :: args))) |
|
570 |
| ast_of t = simple_ast_of ctxt t |
|
42254 | 571 |
and trans a args = ast_of (trf a ctxt' dummyT args) |
572 |
handle Match => Ast.mk_appl (Ast.Constant a) (map ast_of args); |
|
42245 | 573 |
in ast_of tm end; |
574 |
||
575 |
fun sort_to_ast ctxt trf S = ast_of_termT ctxt trf (term_of_sort S); |
|
42247
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
wenzelm
parents:
42245
diff
changeset
|
576 |
fun typ_to_ast ctxt trf T = ast_of_termT ctxt trf (term_of_typ ctxt T); |
42245 | 577 |
|
578 |
||
579 |
(* term_to_ast *) |
|
580 |
||
52177 | 581 |
local |
582 |
||
583 |
fun mark_aprop tm = |
|
584 |
let |
|
585 |
fun aprop t = Syntax.const "_aprop" $ t; |
|
586 |
||
587 |
fun is_prop Ts t = |
|
52464
36497ee02ed8
clarified printing of type constraints vs. annotations, notably in goal display with its Type_Annotation.ignore_free_types;
wenzelm
parents:
52219
diff
changeset
|
588 |
Type_Annotation.clean (Type_Annotation.fastype_of Ts t) = propT |
36497ee02ed8
clarified printing of type constraints vs. annotations, notably in goal display with its Type_Annotation.ignore_free_types;
wenzelm
parents:
52219
diff
changeset
|
589 |
handle TERM _ => false; |
52177 | 590 |
|
591 |
fun is_term (Const ("Pure.term", _) $ _) = true |
|
592 |
| is_term _ = false; |
|
593 |
||
594 |
fun mark _ (t as Const _) = t |
|
595 |
| mark Ts (t as Const ("_bound", _) $ u) = if is_prop Ts u then aprop t else t |
|
52464
36497ee02ed8
clarified printing of type constraints vs. annotations, notably in goal display with its Type_Annotation.ignore_free_types;
wenzelm
parents:
52219
diff
changeset
|
596 |
| mark Ts (t as Free _) = if is_prop Ts t then aprop t else t |
36497ee02ed8
clarified printing of type constraints vs. annotations, notably in goal display with its Type_Annotation.ignore_free_types;
wenzelm
parents:
52219
diff
changeset
|
597 |
| mark Ts (t as Var _) = if is_prop Ts t then aprop t else t |
52177 | 598 |
| mark Ts (t as Bound _) = if is_prop Ts t then aprop t else t |
599 |
| mark Ts (Abs (x, T, t)) = Abs (x, T, mark (T :: Ts) t) |
|
56243 | 600 |
| mark Ts (t as t1 $ (t2 as Const ("Pure.type", Type ("itself", [T])))) = |
52177 | 601 |
if is_prop Ts t andalso not (is_term t) then Const ("_type_prop", T) $ mark Ts t1 |
602 |
else mark Ts t1 $ mark Ts t2 |
|
603 |
| mark Ts (t as t1 $ t2) = |
|
604 |
(if is_Const (Term.head_of t) orelse not (is_prop Ts t) then I else aprop) |
|
605 |
(mark Ts t1 $ mark Ts t2); |
|
606 |
in mark [] tm end; |
|
607 |
||
62239 | 608 |
fun prune_types tm = |
42245 | 609 |
let |
52211
66bc827e37f8
explicit support for type annotations within printed syntax trees;
wenzelm
parents:
52210
diff
changeset
|
610 |
fun regard t t' seen = |
66bc827e37f8
explicit support for type annotations within printed syntax trees;
wenzelm
parents:
52210
diff
changeset
|
611 |
if Type_Annotation.is_omitted (Type_Annotation.fastype_of [] t) then (t, seen) |
66bc827e37f8
explicit support for type annotations within printed syntax trees;
wenzelm
parents:
52210
diff
changeset
|
612 |
else if member (op aconv) seen t then (t', seen) |
66bc827e37f8
explicit support for type annotations within printed syntax trees;
wenzelm
parents:
52210
diff
changeset
|
613 |
else (t, t :: seen); |
42245 | 614 |
|
52211
66bc827e37f8
explicit support for type annotations within printed syntax trees;
wenzelm
parents:
52210
diff
changeset
|
615 |
fun prune (t as Const _, seen) = (t, seen) |
66bc827e37f8
explicit support for type annotations within printed syntax trees;
wenzelm
parents:
52210
diff
changeset
|
616 |
| prune (t as Free (x, T), seen) = regard t (Free (x, Type_Annotation.ignore_type T)) seen |
66bc827e37f8
explicit support for type annotations within printed syntax trees;
wenzelm
parents:
52210
diff
changeset
|
617 |
| prune (t as Var (xi, T), seen) = regard t (Var (xi, Type_Annotation.ignore_type T)) seen |
66bc827e37f8
explicit support for type annotations within printed syntax trees;
wenzelm
parents:
52210
diff
changeset
|
618 |
| prune (t as Bound _, seen) = (t, seen) |
66bc827e37f8
explicit support for type annotations within printed syntax trees;
wenzelm
parents:
52210
diff
changeset
|
619 |
| prune (Abs (x, T, t), seen) = |
52186 | 620 |
let val (t', seen') = prune (t, seen); |
52211
66bc827e37f8
explicit support for type annotations within printed syntax trees;
wenzelm
parents:
52210
diff
changeset
|
621 |
in (Abs (x, T, t'), seen') end |
52186 | 622 |
| prune (t1 $ t2, seen) = |
52177 | 623 |
let |
52186 | 624 |
val (t1', seen') = prune (t1, seen); |
625 |
val (t2', seen'') = prune (t2, seen'); |
|
52177 | 626 |
in (t1' $ t2', seen'') end; |
52186 | 627 |
in #1 (prune (tm, [])) end; |
52177 | 628 |
|
80749 | 629 |
fun mark_atoms ctxt tm = |
52186 | 630 |
let |
59152
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
wenzelm
parents:
59064
diff
changeset
|
631 |
val {structs, fixes} = Syntax_Trans.get_idents ctxt; |
52186 | 632 |
val show_structs = Config.get ctxt show_structs; |
633 |
||
634 |
fun mark ((t as Const (c, _)) $ u) = |
|
52160 | 635 |
if member (op =) Pure_Thy.token_markers c |
52186 | 636 |
then t $ u else mark t $ mark u |
637 |
| mark (t $ u) = mark t $ mark u |
|
638 |
| mark (Abs (x, T, t)) = Abs (x, T, mark t) |
|
639 |
| mark (t as Const (c, T)) = |
|
80749 | 640 |
if Proof_Context.is_syntax_const ctxt c then t |
42245 | 641 |
else Const (Lexicon.mark_const c, T) |
52186 | 642 |
| mark (t as Free (x, T)) = |
42245 | 643 |
let val i = find_index (fn s => s = x) structs + 1 in |
644 |
if i = 0 andalso member (op =) fixes x then |
|
645 |
Const (Lexicon.mark_fixed x, T) |
|
646 |
else if i = 1 andalso not show_structs then |
|
42476 | 647 |
Syntax.const "_struct" $ Syntax.const "_indexdefault" |
648 |
else Syntax.const "_free" $ t |
|
42245 | 649 |
end |
52186 | 650 |
| mark (t as Var (xi, T)) = |
81279 | 651 |
if xi = Auto_Bind.dddot_indexname then Const ("_DDDOT", T) |
42476 | 652 |
else Syntax.const "_var" $ t |
52186 | 653 |
| mark a = a; |
654 |
in mark tm end; |
|
655 |
||
656 |
in |
|
657 |
||
81169 | 658 |
fun term_to_ast ctxt trf = |
52186 | 659 |
let |
660 |
val show_types = Config.get ctxt show_types orelse Config.get ctxt show_sorts; |
|
661 |
val show_markup = Config.get ctxt show_markup; |
|
81166 | 662 |
val show_consts_markup = Config.get ctxt show_consts_markup; |
42245 | 663 |
|
81168 | 664 |
val show_const_types = show_markup andalso show_consts_markup; |
665 |
val show_var_types = show_types orelse show_markup; |
|
666 |
val clean_var_types = show_markup andalso not show_types; |
|
667 |
||
81169 | 668 |
fun constrain clean T ast = |
81196
eb397024b496
suppress dummyT, e.g. from Type_Annotation.print;
wenzelm
parents:
81194
diff
changeset
|
669 |
let val U = Type_Annotation.print clean T in |
eb397024b496
suppress dummyT, e.g. from Type_Annotation.print;
wenzelm
parents:
81194
diff
changeset
|
670 |
if U = dummyT then ast |
81210 | 671 |
else Ast.constrain ast (ast_of_termT ctxt trf (term_of_typ ctxt U)) |
81196
eb397024b496
suppress dummyT, e.g. from Type_Annotation.print;
wenzelm
parents:
81194
diff
changeset
|
672 |
end; |
81164 | 673 |
|
81169 | 674 |
fun main tm = |
42245 | 675 |
(case strip_comb tm of |
81169 | 676 |
(t as Abs _, ts) => Ast.mk_appl (main (Syntax_Trans.abs_tr' ctxt t)) (map main ts) |
42245 | 677 |
| ((c as Const ("_free", _)), Free (x, T) :: ts) => |
81169 | 678 |
Ast.mk_appl (variable (c $ Syntax.free x) T) (map main ts) |
42245 | 679 |
| ((c as Const ("_var", _)), Var (xi, T) :: ts) => |
81169 | 680 |
Ast.mk_appl (variable (c $ Syntax.var xi) T) (map main ts) |
49660
de49d9b4d7bc
more explicit Syntax_Trans.mark_bound_abs/mark_bound_body: preserve type information for show_markup;
wenzelm
parents:
49659
diff
changeset
|
681 |
| ((c as Const ("_bound", B)), Free (x, T) :: ts) => |
de49d9b4d7bc
more explicit Syntax_Trans.mark_bound_abs/mark_bound_body: preserve type information for show_markup;
wenzelm
parents:
49659
diff
changeset
|
682 |
let |
de49d9b4d7bc
more explicit Syntax_Trans.mark_bound_abs/mark_bound_body: preserve type information for show_markup;
wenzelm
parents:
49659
diff
changeset
|
683 |
val X = |
de49d9b4d7bc
more explicit Syntax_Trans.mark_bound_abs/mark_bound_body: preserve type information for show_markup;
wenzelm
parents:
49659
diff
changeset
|
684 |
if show_markup andalso not show_types orelse B <> dummyT then T |
de49d9b4d7bc
more explicit Syntax_Trans.mark_bound_abs/mark_bound_body: preserve type information for show_markup;
wenzelm
parents:
49659
diff
changeset
|
685 |
else dummyT; |
81169 | 686 |
in Ast.mk_appl (variable (c $ Syntax.free x) X) (map main ts) end |
42245 | 687 |
| (Const ("_idtdummy", T), ts) => |
81169 | 688 |
Ast.mk_appl (variable (Syntax.const "_idtdummy") T) (map main ts) |
689 |
| (Const (c, T), ts) => constant c T ts |
|
690 |
| (t, ts) => Ast.mk_appl (simple_ast_of ctxt t) (map main ts)) |
|
42245 | 691 |
|
81169 | 692 |
and constant a T args = |
81166 | 693 |
(case SOME (trf a ctxt (Type_Annotation.smash T) args) handle Match => NONE of |
81169 | 694 |
SOME t => main t |
81166 | 695 |
| NONE => |
81169 | 696 |
let val c = Ast.Constant a |> show_const_types ? constrain {clean = true} T |
697 |
in Ast.mk_appl c (map main args) end) |
|
42245 | 698 |
|
81169 | 699 |
and variable v T = |
81165 | 700 |
simple_ast_of ctxt v |
81169 | 701 |
|> show_var_types ? constrain {clean = clean_var_types} T; |
42245 | 702 |
in |
81169 | 703 |
mark_aprop |
704 |
#> show_types ? prune_types |
|
705 |
#> Variable.revert_bounds ctxt |
|
706 |
#> mark_atoms ctxt |
|
707 |
#> main |
|
42245 | 708 |
end; |
709 |
||
52177 | 710 |
end; |
711 |
||
42245 | 712 |
|
713 |
||
42242
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
714 |
(** unparse **) |
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
715 |
|
42245 | 716 |
local |
717 |
||
81552 | 718 |
fun extern ctxt c = |
719 |
(case Syntax.get_consts (Proof_Context.syntax_of ctxt) c of |
|
720 |
[b] => b |
|
721 |
| bs => |
|
722 |
(case filter Lexicon.is_marked bs of |
|
723 |
[] => c |
|
724 |
| [b] => b |
|
725 |
| _ => error ("Multiple logical entities for " ^ quote c ^ ": " ^ commas_quote bs))) |
|
726 |
|> Proof_Context.extern_entity ctxt; |
|
80743 | 727 |
|
81552 | 728 |
fun pretty_free ctxt x = |
729 |
Pretty.marks_str (markup_free ctxt x, Proof_Context.extern_fixed ctxt x); |
|
730 |
||
731 |
fun pretty_var s = |
|
732 |
(case Lexicon.read_variable s of |
|
733 |
SOME (x, i) => |
|
734 |
(case try Name.dest_skolem x of |
|
735 |
SOME x' => (Markup.skolem, Term.string_of_vname (x', i)) |
|
736 |
| NONE => (Markup.var, s)) |
|
737 |
| NONE => (Markup.var, s)) |
|
738 |
|> Pretty.mark_str; |
|
42267
9566078ad905
simplified printer context: uniform externing and static token translations;
wenzelm
parents:
42264
diff
changeset
|
739 |
|
50201
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents:
49690
diff
changeset
|
740 |
val typing_elem = YXML.output_markup_elem Markup.typing; |
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents:
49690
diff
changeset
|
741 |
val sorting_elem = YXML.output_markup_elem Markup.sorting; |
49657
40e4feac2921
turn constraints into Isabelle_Markup.typing, depending on show_markup options;
wenzelm
parents:
49655
diff
changeset
|
742 |
|
81545
6f8a56a6b391
clarified renaming of bounds, using Syntax_Trans.variant_bounds: avoid structures and fixed variables with syntax;
wenzelm
parents:
81543
diff
changeset
|
743 |
fun exclude_consts ast ctxt = |
42245 | 744 |
let |
81545
6f8a56a6b391
clarified renaming of bounds, using Syntax_Trans.variant_bounds: avoid structures and fixed variables with syntax;
wenzelm
parents:
81543
diff
changeset
|
745 |
val s = the_default "" (Syntax_Trans.default_struct ctxt); |
6f8a56a6b391
clarified renaming of bounds, using Syntax_Trans.variant_bounds: avoid structures and fixed variables with syntax;
wenzelm
parents:
81543
diff
changeset
|
746 |
|
6f8a56a6b391
clarified renaming of bounds, using Syntax_Trans.variant_bounds: avoid structures and fixed variables with syntax;
wenzelm
parents:
81543
diff
changeset
|
747 |
fun exclude (Ast.Appl [Ast.Constant "_struct", Ast.Constant "_indexdefault"]) = Symset.insert s |
6f8a56a6b391
clarified renaming of bounds, using Syntax_Trans.variant_bounds: avoid structures and fixed variables with syntax;
wenzelm
parents:
81543
diff
changeset
|
748 |
| exclude (Ast.Constant c) = |
6f8a56a6b391
clarified renaming of bounds, using Syntax_Trans.variant_bounds: avoid structures and fixed variables with syntax;
wenzelm
parents:
81543
diff
changeset
|
749 |
if Lexicon.is_fixed c then Symset.insert (Lexicon.unmark_fixed c) else I |
6f8a56a6b391
clarified renaming of bounds, using Syntax_Trans.variant_bounds: avoid structures and fixed variables with syntax;
wenzelm
parents:
81543
diff
changeset
|
750 |
| exclude (Ast.Appl [Ast.Constant "_bound", Ast.Variable x]) = Symset.insert x |
81543
fa37ee54644c
clarified printing of consts: rename apart from all bounds, and thus avoid old Term.declare_free_names with its adhoc policy ("as they are printed");
wenzelm
parents:
81540
diff
changeset
|
751 |
| exclude (Ast.Appl [Ast.Constant "_free", Ast.Variable x]) = Symset.insert x |
fa37ee54644c
clarified printing of consts: rename apart from all bounds, and thus avoid old Term.declare_free_names with its adhoc policy ("as they are printed");
wenzelm
parents:
81540
diff
changeset
|
752 |
| exclude (Ast.Appl asts) = fold exclude asts |
fa37ee54644c
clarified printing of consts: rename apart from all bounds, and thus avoid old Term.declare_free_names with its adhoc policy ("as they are printed");
wenzelm
parents:
81540
diff
changeset
|
753 |
| exclude _ = I; |
81545
6f8a56a6b391
clarified renaming of bounds, using Syntax_Trans.variant_bounds: avoid structures and fixed variables with syntax;
wenzelm
parents:
81543
diff
changeset
|
754 |
in Proof_Context.exclude_consts (Symset.build (exclude ast)) ctxt end; |
81543
fa37ee54644c
clarified printing of consts: rename apart from all bounds, and thus avoid old Term.declare_free_names with its adhoc policy ("as they are printed");
wenzelm
parents:
81540
diff
changeset
|
755 |
|
fa37ee54644c
clarified printing of consts: rename apart from all bounds, and thus avoid old Term.declare_free_names with its adhoc policy ("as they are printed");
wenzelm
parents:
81540
diff
changeset
|
756 |
fun unparse_t t_to_ast pretty_flags language_markup ctxt0 t = |
fa37ee54644c
clarified printing of consts: rename apart from all bounds, and thus avoid old Term.declare_free_names with its adhoc policy ("as they are printed");
wenzelm
parents:
81540
diff
changeset
|
757 |
let |
fa37ee54644c
clarified printing of consts: rename apart from all bounds, and thus avoid old Term.declare_free_names with its adhoc policy ("as they are printed");
wenzelm
parents:
81540
diff
changeset
|
758 |
val syn = Proof_Context.syntax_of ctxt0; |
fa37ee54644c
clarified printing of consts: rename apart from all bounds, and thus avoid old Term.declare_free_names with its adhoc policy ("as they are printed");
wenzelm
parents:
81540
diff
changeset
|
759 |
val prtabs = Syntax.print_mode_tabs syn; |
fa37ee54644c
clarified printing of consts: rename apart from all bounds, and thus avoid old Term.declare_free_names with its adhoc policy ("as they are printed");
wenzelm
parents:
81540
diff
changeset
|
760 |
val trf = Syntax.print_ast_translation syn; |
fa37ee54644c
clarified printing of consts: rename apart from all bounds, and thus avoid old Term.declare_free_names with its adhoc policy ("as they are printed");
wenzelm
parents:
81540
diff
changeset
|
761 |
|
fa37ee54644c
clarified printing of consts: rename apart from all bounds, and thus avoid old Term.declare_free_names with its adhoc policy ("as they are printed");
wenzelm
parents:
81540
diff
changeset
|
762 |
val ast = t_to_ast ctxt0 (Syntax.print_translation syn) t; |
fa37ee54644c
clarified printing of consts: rename apart from all bounds, and thus avoid old Term.declare_free_names with its adhoc policy ("as they are printed");
wenzelm
parents:
81540
diff
changeset
|
763 |
val ctxt = exclude_consts ast ctxt0; |
fa37ee54644c
clarified printing of consts: rename apart from all bounds, and thus avoid old Term.declare_free_names with its adhoc policy ("as they are printed");
wenzelm
parents:
81540
diff
changeset
|
764 |
|
49686
678aa92e921c
printed "_ofsort" is subject to show_markup as well;
wenzelm
parents:
49685
diff
changeset
|
765 |
val show_markup = Config.get ctxt show_markup; |
678aa92e921c
printed "_ofsort" is subject to show_markup as well;
wenzelm
parents:
49685
diff
changeset
|
766 |
val show_sorts = Config.get ctxt show_sorts; |
52185 | 767 |
val show_types = Config.get ctxt show_types orelse show_sorts; |
49657
40e4feac2921
turn constraints into Isabelle_Markup.typing, depending on show_markup options;
wenzelm
parents:
49655
diff
changeset
|
768 |
|
81552 | 769 |
val pretty_free_cache = Symtab.apply_unsynchronized_cache (pretty_free ctxt); |
770 |
val pretty_var_cache = Symtab.apply_unsynchronized_cache pretty_var; |
|
81173
6002cb6bfb0a
performance tuning: cache markup and extern operations;
wenzelm
parents:
81169
diff
changeset
|
771 |
|
81558
b57996a0688c
clarified term positions and markup: syntax = true means this is via concrete syntax;
wenzelm
parents:
81552
diff
changeset
|
772 |
val markup_syntax_cache = |
b57996a0688c
clarified term positions and markup: syntax = true means this is via concrete syntax;
wenzelm
parents:
81552
diff
changeset
|
773 |
Symtab.apply_unsynchronized_cache (markup_entity ctxt #> map (Markup.syntax_properties true)); |
81552 | 774 |
|
775 |
val pretty_entity_cache = |
|
776 |
Symtab.apply_unsynchronized_cache (fn a => |
|
81558
b57996a0688c
clarified term positions and markup: syntax = true means this is via concrete syntax;
wenzelm
parents:
81552
diff
changeset
|
777 |
Pretty.marks_str (markup_entity ctxt a, extern ctxt a)); |
49657
40e4feac2921
turn constraints into Isabelle_Markup.typing, depending on show_markup options;
wenzelm
parents:
49655
diff
changeset
|
778 |
|
80882
2cdb00f797b1
performance tuning: cache for highly redundant markup (types and sorts);
wenzelm
parents:
80880
diff
changeset
|
779 |
val cache1 = Unsynchronized.ref (Ast.Table.empty: Markup.output Pretty.block Ast.Table.table); |
2cdb00f797b1
performance tuning: cache for highly redundant markup (types and sorts);
wenzelm
parents:
80880
diff
changeset
|
780 |
val cache2 = Unsynchronized.ref (Ast.Table.empty: Markup.output Pretty.block Ast.Table.table); |
2cdb00f797b1
performance tuning: cache for highly redundant markup (types and sorts);
wenzelm
parents:
80880
diff
changeset
|
781 |
|
50201
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents:
49690
diff
changeset
|
782 |
fun token_trans "_tfree" x = SOME (Pretty.mark_str (Markup.tfree, x)) |
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents:
49690
diff
changeset
|
783 |
| token_trans "_tvar" x = SOME (Pretty.mark_str (Markup.tvar, x)) |
81552 | 784 |
| token_trans "_free" x = SOME (pretty_free_cache x) |
50201
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents:
49690
diff
changeset
|
785 |
| token_trans "_bound" x = SOME (Pretty.mark_str (Markup.bound, x)) |
64677
8dc24130e8fe
more uniform treatment of "bad" like other messages (with serial number);
wenzelm
parents:
63395
diff
changeset
|
786 |
| token_trans "_loose" x = SOME (Pretty.mark_str (Markup.bad (), x)) |
81552 | 787 |
| token_trans "_var" x = SOME (pretty_var_cache x) |
50201
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents:
49690
diff
changeset
|
788 |
| token_trans "_numeral" x = SOME (Pretty.mark_str (Markup.numeral, x)) |
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents:
49690
diff
changeset
|
789 |
| token_trans "_inner_string" x = SOME (Pretty.mark_str (Markup.inner_string, x)) |
42267
9566078ad905
simplified printer context: uniform externing and static token translations;
wenzelm
parents:
42264
diff
changeset
|
790 |
| token_trans _ _ = NONE; |
9566078ad905
simplified printer context: uniform externing and static token translations;
wenzelm
parents:
42264
diff
changeset
|
791 |
|
49655 | 792 |
fun markup_trans a [Ast.Variable x] = token_trans a x |
49662
de6be6922c19
proper handling of constraints stemming from idtyp_ast_tr';
wenzelm
parents:
49660
diff
changeset
|
793 |
| markup_trans "_constrain" [t, ty] = constrain_trans t ty |
de6be6922c19
proper handling of constraints stemming from idtyp_ast_tr';
wenzelm
parents:
49660
diff
changeset
|
794 |
| markup_trans "_idtyp" [t, ty] = constrain_trans t ty |
49686
678aa92e921c
printed "_ofsort" is subject to show_markup as well;
wenzelm
parents:
49685
diff
changeset
|
795 |
| markup_trans "_ofsort" [ty, s] = ofsort_trans ty s |
49657
40e4feac2921
turn constraints into Isabelle_Markup.typing, depending on show_markup options;
wenzelm
parents:
49655
diff
changeset
|
796 |
| markup_trans _ _ = NONE |
49655 | 797 |
|
49662
de6be6922c19
proper handling of constraints stemming from idtyp_ast_tr';
wenzelm
parents:
49660
diff
changeset
|
798 |
and constrain_trans t ty = |
80879 | 799 |
if show_markup andalso not show_types |
80880 | 800 |
then SOME (markup_ast true t ty) else NONE |
49662
de6be6922c19
proper handling of constraints stemming from idtyp_ast_tr';
wenzelm
parents:
49660
diff
changeset
|
801 |
|
49686
678aa92e921c
printed "_ofsort" is subject to show_markup as well;
wenzelm
parents:
49685
diff
changeset
|
802 |
and ofsort_trans ty s = |
80879 | 803 |
if show_markup andalso not show_sorts |
80880 | 804 |
then SOME (markup_ast false ty s) else NONE |
49686
678aa92e921c
printed "_ofsort" is subject to show_markup as well;
wenzelm
parents:
49685
diff
changeset
|
805 |
|
81198 | 806 |
and pretty_ast flags m = |
807 |
Printer.pretty flags ctxt prtabs |
|
81200 | 808 |
{trf = trf, constrain_block = constrain_block true, constrain_trans = constrain_trans, |
81552 | 809 |
markup_trans = markup_trans, markup_syntax = markup_syntax_cache, |
810 |
pretty_entity = pretty_entity_cache} |
|
81198 | 811 |
#> Pretty.markup m |
80848
df85df6315af
clarified signature: prefer explicit type Bytes.T;
wenzelm
parents:
80789
diff
changeset
|
812 |
|
80880 | 813 |
and markup_ast is_typing a A = |
81199 | 814 |
Pretty.make_block (constrain_block is_typing A) |
815 |
[pretty_ast (if is_typing then pretty_flags else Printer.type_mode_flags) Markup.empty a] |
|
80882
2cdb00f797b1
performance tuning: cache for highly redundant markup (types and sorts);
wenzelm
parents:
80880
diff
changeset
|
816 |
|
81199 | 817 |
and constrain_block is_typing A = |
818 |
let val cache = if is_typing then cache1 else cache2 in |
|
819 |
(case Ast.Table.lookup (! cache) A of |
|
820 |
SOME block => block |
|
821 |
| NONE => |
|
822 |
let |
|
823 |
val ((bg1, bg2), en) = if is_typing then typing_elem else sorting_elem; |
|
824 |
val B = Pretty.symbolic_output (pretty_ast Printer.type_mode_flags Markup.empty A); |
|
825 |
val bg = implode (bg1 :: Bytes.contents B @ [bg2]); |
|
826 |
val block = {markup = (bg, en), open_block = false, consistent = false, indent = 0}; |
|
827 |
in Unsynchronized.change cache (Ast.Table.update (A, block)); block end) |
|
828 |
end; |
|
42245 | 829 |
in |
81543
fa37ee54644c
clarified printing of consts: rename apart from all bounds, and thus avoid old Term.declare_free_names with its adhoc policy ("as they are printed");
wenzelm
parents:
81540
diff
changeset
|
830 |
Ast.normalize ctxt {permissive_constraints = true} (Syntax.print_rules syn) ast |
81198 | 831 |
|> pretty_ast pretty_flags language_markup |
42245 | 832 |
end; |
833 |
||
834 |
in |
|
835 |
||
81198 | 836 |
val unparse_sort = unparse_t sort_to_ast Printer.type_mode_flags (Markup.language_sort false); |
837 |
val unparse_typ = unparse_t typ_to_ast Printer.type_mode_flags (Markup.language_type false); |
|
42241
dd8029f71e1c
separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff
changeset
|
838 |
|
dd8029f71e1c
separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff
changeset
|
839 |
fun unparse_term ctxt = |
dd8029f71e1c
separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff
changeset
|
840 |
let |
42360 | 841 |
val thy = Proof_Context.theory_of ctxt; |
81198 | 842 |
val pretty_flags = {type_mode = false, curried = not (Pure_Thy.old_appl_syntax thy)}; |
843 |
in unparse_t term_to_ast pretty_flags (Markup.language_term false) ctxt end; |
|
42241
dd8029f71e1c
separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff
changeset
|
844 |
|
42249 | 845 |
end; |
846 |
||
42241
dd8029f71e1c
separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff
changeset
|
847 |
|
42245 | 848 |
|
849 |
(** translations **) |
|
850 |
||
851 |
(* type propositions *) |
|
852 |
||
62529
8b7bdfc09f3b
clarified treatment of fragments of Isabelle symbols during bootstrap;
wenzelm
parents:
62505
diff
changeset
|
853 |
fun type_prop_tr' ctxt T [Const ("\<^const>Pure.sort_constraint", _)] = |
42476 | 854 |
Syntax.const "_sort_constraint" $ term_of_typ (Config.put show_sorts true ctxt) T |
42247
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
wenzelm
parents:
42245
diff
changeset
|
855 |
| type_prop_tr' ctxt T [t] = |
42476 | 856 |
Syntax.const "_ofclass" $ term_of_typ ctxt T $ t |
42245 | 857 |
| type_prop_tr' _ T ts = raise TYPE ("type_prop_tr'", [T], ts); |
858 |
||
859 |
||
860 |
(* type reflection *) |
|
861 |
||
81201 | 862 |
fun type_tr' ctxt (ty as Type ("itself", [T])) ts = |
863 |
Term.list_comb (Const ("_TYPE", ty) $ term_of_typ ctxt T, ts) |
|
42245 | 864 |
| type_tr' _ _ _ = raise Match; |
865 |
||
866 |
||
867 |
(* type constraints *) |
|
868 |
||
42247
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
wenzelm
parents:
42245
diff
changeset
|
869 |
fun type_constraint_tr' ctxt (Type ("fun", [T, _])) (t :: ts) = |
42476 | 870 |
Term.list_comb (Syntax.const "_constrain" $ t $ term_of_typ ctxt T, ts) |
42245 | 871 |
| type_constraint_tr' _ _ _ = raise Match; |
872 |
||
873 |
||
42295
8fdbb3b10beb
moved CONST syntax/translations to their proper place;
wenzelm
parents:
42294
diff
changeset
|
874 |
(* authentic syntax *) |
8fdbb3b10beb
moved CONST syntax/translations to their proper place;
wenzelm
parents:
42294
diff
changeset
|
875 |
|
55960
beef468837b1
proper position for decode_pos, which is relevant for completion;
wenzelm
parents:
55959
diff
changeset
|
876 |
fun const_ast_tr intern ctxt asts = |
beef468837b1
proper position for decode_pos, which is relevant for completion;
wenzelm
parents:
55959
diff
changeset
|
877 |
(case asts of |
beef468837b1
proper position for decode_pos, which is relevant for completion;
wenzelm
parents:
55959
diff
changeset
|
878 |
[Ast.Appl [Ast.Constant "_constrain", Ast.Variable c, T as Ast.Variable p]] => |
42295
8fdbb3b10beb
moved CONST syntax/translations to their proper place;
wenzelm
parents:
42294
diff
changeset
|
879 |
let |
81558
b57996a0688c
clarified term positions and markup: syntax = true means this is via concrete syntax;
wenzelm
parents:
81552
diff
changeset
|
880 |
val (c', _) = decode_const ctxt (c, map #pos (Term_Position.decode p)); |
42295
8fdbb3b10beb
moved CONST syntax/translations to their proper place;
wenzelm
parents:
42294
diff
changeset
|
881 |
val d = if intern then Lexicon.mark_const c' else c; |
81210 | 882 |
in Ast.constrain (Ast.Constant d) T end |
55960
beef468837b1
proper position for decode_pos, which is relevant for completion;
wenzelm
parents:
55959
diff
changeset
|
883 |
| _ => raise Ast.AST ("const_ast_tr", asts)); |
42295
8fdbb3b10beb
moved CONST syntax/translations to their proper place;
wenzelm
parents:
42294
diff
changeset
|
884 |
|
8fdbb3b10beb
moved CONST syntax/translations to their proper place;
wenzelm
parents:
42294
diff
changeset
|
885 |
|
42245 | 886 |
(* setup translations *) |
887 |
||
53171 | 888 |
val _ = Theory.setup |
52143 | 889 |
(Sign.parse_ast_translation |
890 |
[("_context_const", const_ast_tr true), |
|
891 |
("_context_xconst", const_ast_tr false)] #> |
|
892 |
Sign.typed_print_translation |
|
42245 | 893 |
[("_type_prop", type_prop_tr'), |
62529
8b7bdfc09f3b
clarified treatment of fragments of Isabelle symbols during bootstrap;
wenzelm
parents:
62505
diff
changeset
|
894 |
("\<^const>Pure.type", type_tr'), |
53171 | 895 |
("_type_constraint_", type_constraint_tr')]); |
42245 | 896 |
|
897 |
||
898 |
||
42382
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents:
42379
diff
changeset
|
899 |
(** check/uncheck **) |
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents:
42379
diff
changeset
|
900 |
|
45429 | 901 |
(* context-sensitive (un)checking *) |
902 |
||
903 |
type key = int * bool; |
|
904 |
||
905 |
structure Checks = Generic_Data |
|
906 |
( |
|
907 |
type 'a check = 'a list -> Proof.context -> ('a list * Proof.context) option; |
|
908 |
type T = |
|
909 |
((key * ((string * typ check) * stamp) list) list * |
|
910 |
(key * ((string * term check) * stamp) list) list); |
|
911 |
val empty = ([], []); |
|
912 |
fun merge ((typ_checks1, term_checks1), (typ_checks2, term_checks2)) : T = |
|
913 |
(AList.join (op =) (K (Library.merge (eq_snd (op =)))) (typ_checks1, typ_checks2), |
|
914 |
AList.join (op =) (K (Library.merge (eq_snd (op =)))) (term_checks1, term_checks2)); |
|
915 |
); |
|
916 |
||
917 |
fun print_checks ctxt = |
|
918 |
let |
|
919 |
fun split_checks checks = |
|
920 |
List.partition (fn ((_, un), _) => not un) checks |
|
59058
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents:
58978
diff
changeset
|
921 |
|> apply2 (map (fn ((i, _), fs) => (i, map (fst o fst) fs)) |
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents:
58978
diff
changeset
|
922 |
#> sort (int_ord o apply2 fst)); |
45429 | 923 |
fun pretty_checks kind checks = |
924 |
checks |> map (fn (i, names) => Pretty.block |
|
925 |
[Pretty.str (kind ^ " (stage " ^ signed_string_of_int i ^ "):"), |
|
926 |
Pretty.brk 1, Pretty.strs names]); |
|
927 |
||
928 |
val (typs, terms) = Checks.get (Context.Proof ctxt); |
|
929 |
val (typ_checks, typ_unchecks) = split_checks typs; |
|
930 |
val (term_checks, term_unchecks) = split_checks terms; |
|
931 |
in |
|
932 |
pretty_checks "typ_checks" typ_checks @ |
|
933 |
pretty_checks "term_checks" term_checks @ |
|
934 |
pretty_checks "typ_unchecks" typ_unchecks @ |
|
935 |
pretty_checks "term_unchecks" term_unchecks |
|
56334
6b3739fee456
some shortcuts for chunks, which sometimes avoid bulky string output;
wenzelm
parents:
56333
diff
changeset
|
936 |
end |> Pretty.writeln_chunks; |
45429 | 937 |
|
938 |
||
939 |
local |
|
940 |
||
941 |
fun context_check which (key: key) name f = |
|
942 |
Checks.map (which (AList.map_default op = (key, []) (cons ((name, f), stamp ())))); |
|
943 |
||
944 |
fun simple_check eq f xs ctxt = |
|
945 |
let val xs' = f ctxt xs |
|
946 |
in if eq_list eq (xs, xs') then NONE else SOME (xs', ctxt) end; |
|
947 |
||
948 |
in |
|
949 |
||
45444 | 950 |
fun typ_check' stage = context_check apfst (stage, false); |
951 |
fun term_check' stage = context_check apsnd (stage, false); |
|
952 |
fun typ_uncheck' stage = context_check apfst (stage, true); |
|
953 |
fun term_uncheck' stage = context_check apsnd (stage, true); |
|
45429 | 954 |
|
45444 | 955 |
fun typ_check key name f = typ_check' key name (simple_check (op =) f); |
956 |
fun term_check key name f = term_check' key name (simple_check (op aconv) f); |
|
957 |
fun typ_uncheck key name f = typ_uncheck' key name (simple_check (op =) f); |
|
958 |
fun term_uncheck key name f = term_uncheck' key name (simple_check (op aconv) f); |
|
45429 | 959 |
|
960 |
end; |
|
961 |
||
962 |
||
963 |
local |
|
964 |
||
965 |
fun check_stage fs = perhaps_loop (perhaps_apply (map uncurry fs)); |
|
966 |
fun check_all fs = perhaps_apply (map check_stage fs); |
|
967 |
||
968 |
fun check which uncheck ctxt0 xs0 = |
|
969 |
let |
|
970 |
val funs = which (Checks.get (Context.Proof ctxt0)) |
|
971 |
|> map_filter (fn ((i, u), fs) => if uncheck = u then SOME (i, map (snd o fst) fs) else NONE) |
|
59058
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents:
58978
diff
changeset
|
972 |
|> Library.sort (int_ord o apply2 fst) |> map snd |
45429 | 973 |
|> not uncheck ? map rev; |
974 |
in #1 (perhaps (check_all funs) (xs0, ctxt0)) end; |
|
975 |
||
976 |
val apply_typ_check = check fst false; |
|
977 |
val apply_term_check = check snd false; |
|
978 |
val apply_typ_uncheck = check fst true; |
|
979 |
val apply_term_uncheck = check snd true; |
|
980 |
||
981 |
in |
|
982 |
||
49674
dbadb4d03cbc
report sort assignment of visible type variables;
wenzelm
parents:
49662
diff
changeset
|
983 |
fun check_typs ctxt raw_tys = |
dbadb4d03cbc
report sort assignment of visible type variables;
wenzelm
parents:
49662
diff
changeset
|
984 |
let |
dbadb4d03cbc
report sort assignment of visible type variables;
wenzelm
parents:
49662
diff
changeset
|
985 |
val (sorting_report, tys) = Proof_Context.prepare_sortsT ctxt raw_tys; |
71675 | 986 |
val _ = if Context_Position.reports_enabled ctxt then Output.report sorting_report else (); |
49674
dbadb4d03cbc
report sort assignment of visible type variables;
wenzelm
parents:
49662
diff
changeset
|
987 |
in |
dbadb4d03cbc
report sort assignment of visible type variables;
wenzelm
parents:
49662
diff
changeset
|
988 |
tys |
dbadb4d03cbc
report sort assignment of visible type variables;
wenzelm
parents:
49662
diff
changeset
|
989 |
|> apply_typ_check ctxt |
dbadb4d03cbc
report sort assignment of visible type variables;
wenzelm
parents:
49662
diff
changeset
|
990 |
|> Term_Sharing.typs (Proof_Context.theory_of ctxt) |
dbadb4d03cbc
report sort assignment of visible type variables;
wenzelm
parents:
49662
diff
changeset
|
991 |
end; |
45423
92f91f990165
tuned signature -- emphasize internal role of these operations;
wenzelm
parents:
45412
diff
changeset
|
992 |
|
45445
41e641a870de
pass term positions into check phase, where resulting types are reported accordingly, and eventually shown as tooltips;
wenzelm
parents:
45444
diff
changeset
|
993 |
fun check_terms ctxt raw_ts = |
41e641a870de
pass term positions into check phase, where resulting types are reported accordingly, and eventually shown as tooltips;
wenzelm
parents:
45444
diff
changeset
|
994 |
let |
49674
dbadb4d03cbc
report sort assignment of visible type variables;
wenzelm
parents:
49662
diff
changeset
|
995 |
val (sorting_report, raw_ts') = Proof_Context.prepare_sorts ctxt raw_ts; |
dbadb4d03cbc
report sort assignment of visible type variables;
wenzelm
parents:
49662
diff
changeset
|
996 |
val (ts, ps) = Type_Infer_Context.prepare_positions ctxt raw_ts'; |
dbadb4d03cbc
report sort assignment of visible type variables;
wenzelm
parents:
49662
diff
changeset
|
997 |
|
45445
41e641a870de
pass term positions into check phase, where resulting types are reported accordingly, and eventually shown as tooltips;
wenzelm
parents:
45444
diff
changeset
|
998 |
val tys = map (Logic.mk_type o snd) ps; |
45448
018f8959c7a6
more efficient prepare_sorts -- bypass encoded positions;
wenzelm
parents:
45447
diff
changeset
|
999 |
val (ts', tys') = ts @ tys |
45445
41e641a870de
pass term positions into check phase, where resulting types are reported accordingly, and eventually shown as tooltips;
wenzelm
parents:
45444
diff
changeset
|
1000 |
|> apply_term_check ctxt |
41e641a870de
pass term positions into check phase, where resulting types are reported accordingly, and eventually shown as tooltips;
wenzelm
parents:
45444
diff
changeset
|
1001 |
|> chop (length ts); |
49674
dbadb4d03cbc
report sort assignment of visible type variables;
wenzelm
parents:
49662
diff
changeset
|
1002 |
val typing_report = |
81558
b57996a0688c
clarified term positions and markup: syntax = true means this is via concrete syntax;
wenzelm
parents:
81552
diff
changeset
|
1003 |
fold2 (fn ({pos, ...}, _) => fn ty => |
45445
41e641a870de
pass term positions into check phase, where resulting types are reported accordingly, and eventually shown as tooltips;
wenzelm
parents:
45444
diff
changeset
|
1004 |
if Position.is_reported pos then |
50201
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents:
49690
diff
changeset
|
1005 |
cons (Position.reported_text pos Markup.typing |
49674
dbadb4d03cbc
report sort assignment of visible type variables;
wenzelm
parents:
49662
diff
changeset
|
1006 |
(Syntax.string_of_typ ctxt (Logic.dest_type ty))) |
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:
56294
diff
changeset
|
1007 |
else I) ps tys' []; |
49674
dbadb4d03cbc
report sort assignment of visible type variables;
wenzelm
parents:
49662
diff
changeset
|
1008 |
|
56294
85911b8a6868
prefer Context_Position where a context is available;
wenzelm
parents:
56243
diff
changeset
|
1009 |
val _ = |
71675 | 1010 |
if Context_Position.reports_enabled ctxt |
1011 |
then Output.report (sorting_report @ typing_report) else (); |
|
45445
41e641a870de
pass term positions into check phase, where resulting types are reported accordingly, and eventually shown as tooltips;
wenzelm
parents:
45444
diff
changeset
|
1012 |
in Term_Sharing.terms (Proof_Context.theory_of ctxt) ts' end; |
45423
92f91f990165
tuned signature -- emphasize internal role of these operations;
wenzelm
parents:
45412
diff
changeset
|
1013 |
|
42382
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents:
42379
diff
changeset
|
1014 |
fun check_props ctxt = map (Type.constraint propT) #> check_terms ctxt; |
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents:
42379
diff
changeset
|
1015 |
|
45429 | 1016 |
val uncheck_typs = apply_typ_uncheck; |
1017 |
val uncheck_terms = apply_term_uncheck; |
|
1018 |
||
1019 |
end; |
|
1020 |
||
1021 |
||
62952 | 1022 |
(* install operations *) |
42242
39261908e12f
moved decode/parse operations to standard_syntax.ML;
wenzelm
parents:
42241
diff
changeset
|
1023 |
|
62897
8093203f0b89
prefer regular context update, to allow continuous editing of Pure;
wenzelm
parents:
62800
diff
changeset
|
1024 |
val _ = |
8093203f0b89
prefer regular context update, to allow continuous editing of Pure;
wenzelm
parents:
62800
diff
changeset
|
1025 |
Theory.setup |
8093203f0b89
prefer regular context update, to allow continuous editing of Pure;
wenzelm
parents:
62800
diff
changeset
|
1026 |
(Syntax.install_operations |
8093203f0b89
prefer regular context update, to allow continuous editing of Pure;
wenzelm
parents:
62800
diff
changeset
|
1027 |
{parse_sort = parse_sort, |
8093203f0b89
prefer regular context update, to allow continuous editing of Pure;
wenzelm
parents:
62800
diff
changeset
|
1028 |
parse_typ = parse_typ, |
8093203f0b89
prefer regular context update, to allow continuous editing of Pure;
wenzelm
parents:
62800
diff
changeset
|
1029 |
parse_term = parse_term false, |
8093203f0b89
prefer regular context update, to allow continuous editing of Pure;
wenzelm
parents:
62800
diff
changeset
|
1030 |
parse_prop = parse_term true, |
8093203f0b89
prefer regular context update, to allow continuous editing of Pure;
wenzelm
parents:
62800
diff
changeset
|
1031 |
unparse_sort = unparse_sort, |
8093203f0b89
prefer regular context update, to allow continuous editing of Pure;
wenzelm
parents:
62800
diff
changeset
|
1032 |
unparse_typ = unparse_typ, |
8093203f0b89
prefer regular context update, to allow continuous editing of Pure;
wenzelm
parents:
62800
diff
changeset
|
1033 |
unparse_term = unparse_term, |
8093203f0b89
prefer regular context update, to allow continuous editing of Pure;
wenzelm
parents:
62800
diff
changeset
|
1034 |
check_typs = check_typs, |
8093203f0b89
prefer regular context update, to allow continuous editing of Pure;
wenzelm
parents:
62800
diff
changeset
|
1035 |
check_terms = check_terms, |
8093203f0b89
prefer regular context update, to allow continuous editing of Pure;
wenzelm
parents:
62800
diff
changeset
|
1036 |
check_props = check_props, |
8093203f0b89
prefer regular context update, to allow continuous editing of Pure;
wenzelm
parents:
62800
diff
changeset
|
1037 |
uncheck_typs = uncheck_typs, |
8093203f0b89
prefer regular context update, to allow continuous editing of Pure;
wenzelm
parents:
62800
diff
changeset
|
1038 |
uncheck_terms = uncheck_terms}); |
42241
dd8029f71e1c
separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff
changeset
|
1039 |
|
dd8029f71e1c
separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff
changeset
|
1040 |
end; |
62952 | 1041 |
|
1042 |
||
1043 |
(* standard phases *) |
|
1044 |
||
1045 |
val _ = Context.>> |
|
1046 |
(Syntax_Phases.typ_check 0 "standard" Proof_Context.standard_typ_check #> |
|
1047 |
Syntax_Phases.term_check 0 "standard" |
|
1048 |
(fn ctxt => Type_Infer_Context.infer_types ctxt #> map (Proof_Context.expand_abbrevs ctxt)) #> |
|
1049 |
Syntax_Phases.term_check 100 "standard_finish" Proof_Context.standard_term_check_finish #> |
|
1050 |
Syntax_Phases.term_uncheck 0 "standard" Proof_Context.standard_term_uncheck); |