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