| author | bulwahn | 
| Fri, 21 Oct 2011 14:25:38 +0200 | |
| changeset 45236 | ac4a2a66707d | 
| parent 44736 | c2a3f1c84179 | 
| child 45412 | 7797f5351ec4 | 
| 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 | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 10 | val term_sorts: term -> (indexname * sort) list | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 11 | val typ_of_term: (indexname -> sort) -> term -> typ | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 12 | val decode_term: Proof.context -> | 
| 44736 | 13 | Position.report list * term Exn.result -> Position.report list * term Exn.result | 
| 42242 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 14 | val parse_ast_pattern: Proof.context -> string * string -> Ast.ast | 
| 42247 
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
 wenzelm parents: 
42245diff
changeset | 15 | val term_of_typ: Proof.context -> typ -> term | 
| 42242 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 16 | end | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 17 | |
| 42243 | 18 | structure Syntax_Phases: SYNTAX_PHASES = | 
| 42241 
dd8029f71e1c
separate module for standard implementation of inner syntax operations;
 wenzelm parents: diff
changeset | 19 | struct | 
| 
dd8029f71e1c
separate module for standard implementation of inner syntax operations;
 wenzelm parents: diff
changeset | 20 | |
| 42298 
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
 wenzelm parents: 
42296diff
changeset | 21 | (** markup logical entities **) | 
| 
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
 wenzelm parents: 
42296diff
changeset | 22 | |
| 
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
 wenzelm parents: 
42296diff
changeset | 23 | fun markup_class ctxt c = | 
| 42379 | 24 | [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 | 25 | |
| 
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
 wenzelm parents: 
42296diff
changeset | 26 | fun markup_type ctxt c = | 
| 42379 | 27 | [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 | 28 | |
| 
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
 wenzelm parents: 
42296diff
changeset | 29 | fun markup_const ctxt c = | 
| 42379 | 30 | [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 | 31 | |
| 
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
 wenzelm parents: 
42296diff
changeset | 32 | fun markup_free ctxt x = | 
| 
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
 wenzelm parents: 
42296diff
changeset | 33 | [if can Name.dest_skolem x then Markup.skolem else Markup.free] @ | 
| 42493 
01430341fc79
more informative markup for fixed variables (via name space entry);
 wenzelm parents: 
42488diff
changeset | 34 | (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 | 35 | then [Variable.markup_fixed ctxt x] | 
| 
01430341fc79
more informative markup for fixed variables (via name space entry);
 wenzelm parents: 
42488diff
changeset | 36 | else []); | 
| 42298 
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
 wenzelm parents: 
42296diff
changeset | 37 | |
| 
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
 wenzelm parents: 
42296diff
changeset | 38 | fun markup_var xi = [Markup.name (Term.string_of_vname xi) Markup.var]; | 
| 
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
 wenzelm parents: 
42296diff
changeset | 39 | |
| 
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
 wenzelm parents: 
42296diff
changeset | 40 | fun markup_bound def id = | 
| 
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
 wenzelm parents: 
42296diff
changeset | 41 | [Markup.properties [(if def then Markup.defN else Markup.refN, id)] Markup.bound]; | 
| 
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
 wenzelm parents: 
42296diff
changeset | 42 | |
| 
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
 wenzelm parents: 
42296diff
changeset | 43 | fun markup_entity ctxt c = | 
| 42360 | 44 | (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 | 45 | SOME "" => [] | 
| 
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
 wenzelm parents: 
42296diff
changeset | 46 | | 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 | 47 | | 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 | 48 |      {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 | 49 | 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 | 50 | 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 | 51 | 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 | 52 | case_default = K []}); | 
| 
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 | |
| 
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
 wenzelm parents: 
42296diff
changeset | 55 | |
| 42242 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 56 | (** decode parse trees **) | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 57 | |
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 58 | (* sort_of_term *) | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 59 | |
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 60 | fun sort_of_term tm = | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 61 | let | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 62 |     fun err () = raise TERM ("sort_of_term: bad encoding of classes", [tm]);
 | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 63 | |
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 64 | fun class s = Lexicon.unmark_class s handle Fail _ => err (); | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 65 | |
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 66 | fun classes (Const (s, _)) = [class s] | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 67 |       | classes (Const ("_classes", _) $ Const (s, _) $ cs) = class s :: classes cs
 | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 68 | | classes _ = err (); | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 69 | |
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 70 |     fun sort (Const ("_topsort", _)) = []
 | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 71 | | sort (Const (s, _)) = [class s] | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 72 |       | sort (Const ("_sort", _) $ cs) = classes cs
 | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 73 | | sort _ = err (); | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 74 | in sort tm end; | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 75 | |
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 76 | |
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 77 | (* term_sorts *) | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 78 | |
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 79 | fun term_sorts tm = | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 80 | let | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 81 | val sort_of = sort_of_term; | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 82 | |
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 83 |     fun add_env (Const ("_ofsort", _) $ Free (x, _) $ cs) =
 | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 84 | insert (op =) ((x, ~1), sort_of cs) | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 85 |       | add_env (Const ("_ofsort", _) $ (Const ("_tfree", _) $ Free (x, _)) $ cs) =
 | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 86 | insert (op =) ((x, ~1), sort_of cs) | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 87 |       | add_env (Const ("_ofsort", _) $ Var (xi, _) $ cs) =
 | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 88 | insert (op =) (xi, sort_of cs) | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 89 |       | add_env (Const ("_ofsort", _) $ (Const ("_tvar", _) $ Var (xi, _)) $ cs) =
 | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 90 | insert (op =) (xi, sort_of cs) | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 91 | | add_env (Abs (_, _, t)) = add_env t | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 92 | | add_env (t1 $ t2) = add_env t1 #> add_env t2 | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 93 | | add_env _ = I; | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 94 | in add_env tm [] end; | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 95 | |
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 96 | |
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 97 | (* typ_of_term *) | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 98 | |
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 99 | fun typ_of_term get_sort tm = | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 100 | let | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 101 |     fun err () = raise TERM ("typ_of_term: bad encoding of type", [tm]);
 | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 102 | |
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 103 | fun typ_of (Free (x, _)) = TFree (x, get_sort (x, ~1)) | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 104 | | typ_of (Var (xi, _)) = TVar (xi, get_sort xi) | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 105 |       | typ_of (Const ("_tfree",_) $ (t as Free _)) = typ_of t
 | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 106 |       | typ_of (Const ("_tvar",_) $ (t as Var _)) = typ_of t
 | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 107 |       | typ_of (Const ("_ofsort", _) $ Free (x, _) $ _) = TFree (x, get_sort (x, ~1))
 | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 108 |       | typ_of (Const ("_ofsort", _) $ (Const ("_tfree",_) $ Free (x, _)) $ _) =
 | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 109 | TFree (x, get_sort (x, ~1)) | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 110 |       | typ_of (Const ("_ofsort", _) $ Var (xi, _) $ _) = TVar (xi, get_sort xi)
 | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 111 |       | typ_of (Const ("_ofsort", _) $ (Const ("_tvar",_) $ Var (xi, _)) $ _) =
 | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 112 | TVar (xi, get_sort xi) | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 113 |       | typ_of (Const ("_dummy_ofsort", _) $ t) = TFree ("'_dummy_", sort_of_term t)
 | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 114 | | typ_of t = | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 115 | let | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 116 | val (head, args) = Term.strip_comb t; | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 117 | val a = | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 118 | (case head of | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 119 | Const (c, _) => (Lexicon.unmark_type c handle Fail _ => err ()) | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 120 | | _ => err ()); | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 121 | in Type (a, map typ_of args) end; | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 122 | in typ_of tm end; | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 123 | |
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 124 | |
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 125 | (* parsetree_to_ast *) | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 126 | |
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 127 | fun parsetree_to_ast ctxt constrain_pos trf parsetree = | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 128 | let | 
| 44736 | 129 | val reports = Unsynchronized.ref ([]: Position.report list); | 
| 44735 | 130 | fun report pos = Position.store_reports reports [pos]; | 
| 42242 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 131 | |
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 132 | fun trans a args = | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 133 | (case trf a of | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 134 | NONE => Ast.mk_appl (Ast.Constant a) args | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 135 | | SOME f => f ctxt args); | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 136 | |
| 42282 
4fa41e068ff0
report literal tokens according to parsetree head;
 wenzelm parents: 
42281diff
changeset | 137 |     fun asts_of (Parser.Node ("_class_name", [Parser.Tip tok])) =
 | 
| 42242 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 138 | let | 
| 42470 | 139 | val pos = Lexicon.pos_of_token tok; | 
| 140 | val c = Proof_Context.read_class ctxt (Lexicon.str_of_token tok) | |
| 141 | handle ERROR msg => error (msg ^ Position.str_of pos); | |
| 142 | val _ = report pos (markup_class ctxt) c; | |
| 42282 
4fa41e068ff0
report literal tokens according to parsetree head;
 wenzelm parents: 
42281diff
changeset | 143 | in [Ast.Constant (Lexicon.mark_class c)] end | 
| 
4fa41e068ff0
report literal tokens according to parsetree head;
 wenzelm parents: 
42281diff
changeset | 144 |       | asts_of (Parser.Node ("_type_name", [Parser.Tip tok])) =
 | 
| 42242 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 145 | let | 
| 42470 | 146 | val pos = Lexicon.pos_of_token tok; | 
| 147 | val Type (c, _) = | |
| 148 | Proof_Context.read_type_name_proper ctxt false (Lexicon.str_of_token tok) | |
| 149 | handle ERROR msg => error (msg ^ Position.str_of pos); | |
| 150 | val _ = report pos (markup_type ctxt) c; | |
| 42282 
4fa41e068ff0
report literal tokens according to parsetree head;
 wenzelm parents: 
42281diff
changeset | 151 | in [Ast.Constant (Lexicon.mark_type c)] end | 
| 42410 
16bc5564535f
less bulky "_position", for improved readability of parse trees;
 wenzelm parents: 
42408diff
changeset | 152 |       | asts_of (Parser.Node ("_position", [pt as Parser.Tip tok])) =
 | 
| 42242 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 153 | if constrain_pos then | 
| 42282 
4fa41e068ff0
report literal tokens according to parsetree head;
 wenzelm parents: 
42281diff
changeset | 154 | [Ast.Appl [Ast.Constant "_constrain", ast_of pt, | 
| 
4fa41e068ff0
report literal tokens according to parsetree head;
 wenzelm parents: 
42281diff
changeset | 155 | Ast.Variable (Term_Position.encode (Lexicon.pos_of_token tok))]] | 
| 
4fa41e068ff0
report literal tokens according to parsetree head;
 wenzelm parents: 
42281diff
changeset | 156 | else [ast_of pt] | 
| 
4fa41e068ff0
report literal tokens according to parsetree head;
 wenzelm parents: 
42281diff
changeset | 157 | | asts_of (Parser.Node (a, pts)) = | 
| 
4fa41e068ff0
report literal tokens according to parsetree head;
 wenzelm parents: 
42281diff
changeset | 158 | let | 
| 
4fa41e068ff0
report literal tokens according to parsetree head;
 wenzelm parents: 
42281diff
changeset | 159 | val _ = pts |> List.app | 
| 
4fa41e068ff0
report literal tokens according to parsetree head;
 wenzelm parents: 
42281diff
changeset | 160 | (fn Parser.Node _ => () | Parser.Tip tok => | 
| 
4fa41e068ff0
report literal tokens according to parsetree head;
 wenzelm parents: 
42281diff
changeset | 161 | if Lexicon.valued_token tok then () | 
| 42298 
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
 wenzelm parents: 
42296diff
changeset | 162 | else report (Lexicon.pos_of_token tok) (markup_entity ctxt) a); | 
| 42282 
4fa41e068ff0
report literal tokens according to parsetree head;
 wenzelm parents: 
42281diff
changeset | 163 | in [trans a (maps asts_of pts)] end | 
| 
4fa41e068ff0
report literal tokens according to parsetree head;
 wenzelm parents: 
42281diff
changeset | 164 | | asts_of (Parser.Tip tok) = | 
| 
4fa41e068ff0
report literal tokens according to parsetree head;
 wenzelm parents: 
42281diff
changeset | 165 | if Lexicon.valued_token tok | 
| 
4fa41e068ff0
report literal tokens according to parsetree head;
 wenzelm parents: 
42281diff
changeset | 166 | then [Ast.Variable (Lexicon.str_of_token tok)] | 
| 
4fa41e068ff0
report literal tokens according to parsetree head;
 wenzelm parents: 
42281diff
changeset | 167 | else [] | 
| 
4fa41e068ff0
report literal tokens according to parsetree head;
 wenzelm parents: 
42281diff
changeset | 168 | |
| 
4fa41e068ff0
report literal tokens according to parsetree head;
 wenzelm parents: 
42281diff
changeset | 169 | and ast_of pt = | 
| 
4fa41e068ff0
report literal tokens according to parsetree head;
 wenzelm parents: 
42281diff
changeset | 170 | (case asts_of pt of | 
| 
4fa41e068ff0
report literal tokens according to parsetree head;
 wenzelm parents: 
42281diff
changeset | 171 | [ast] => ast | 
| 
4fa41e068ff0
report literal tokens according to parsetree head;
 wenzelm parents: 
42281diff
changeset | 172 |       | asts => raise Ast.AST ("parsetree_to_ast: malformed parsetree", asts));
 | 
| 42242 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 173 | |
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 174 | val ast = Exn.interruptible_capture ast_of parsetree; | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 175 | in (! reports, ast) end; | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 176 | |
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 177 | |
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 178 | (* ast_to_term *) | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 179 | |
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 180 | fun ast_to_term ctxt trf = | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 181 | let | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 182 | fun trans a args = | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 183 | (case trf a of | 
| 42476 | 184 | NONE => Term.list_comb (Syntax.const a, args) | 
| 42242 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 185 | | SOME f => f ctxt args); | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 186 | |
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 187 | fun term_of (Ast.Constant a) = trans a [] | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 188 | | term_of (Ast.Variable x) = Lexicon.read_var x | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 189 | | term_of (Ast.Appl (Ast.Constant a :: (asts as _ :: _))) = | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 190 | trans a (map term_of asts) | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 191 | | term_of (Ast.Appl (ast :: (asts as _ :: _))) = | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 192 | Term.list_comb (term_of ast, map term_of asts) | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 193 |       | 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 | 194 | in term_of end; | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 195 | |
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 196 | |
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 197 | (* decode_term -- transform parse tree into raw term *) | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 198 | |
| 44736 | 199 | fun decode_term _ (result as (_: Position.report list, Exn.Exn _)) = result | 
| 43761 
e72ba84ae58f
tuned signature -- corresponding to Scala version;
 wenzelm parents: 
43731diff
changeset | 200 | | decode_term ctxt (reports0, Exn.Res tm) = | 
| 42242 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 201 | let | 
| 42250 
cc5ac538f991
eliminated odd object-oriented type_context/term_context;
 wenzelm parents: 
42249diff
changeset | 202 | fun get_const a = | 
| 42360 | 203 | ((true, #1 (Term.dest_Const (Proof_Context.read_const_proper ctxt false a))) | 
| 204 | handle ERROR _ => (false, Consts.intern (Proof_Context.consts_of ctxt) a)); | |
| 205 | val get_free = Proof_Context.intern_skolem ctxt; | |
| 42250 
cc5ac538f991
eliminated odd object-oriented type_context/term_context;
 wenzelm parents: 
42249diff
changeset | 206 | |
| 42360 | 207 | val decodeT = typ_of_term (Proof_Context.get_sort ctxt (term_sorts tm)); | 
| 42242 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 208 | |
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 209 | val reports = Unsynchronized.ref reports0; | 
| 44735 | 210 | fun report ps = Position.store_reports reports ps; | 
| 42242 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 211 | |
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 212 |         fun decode ps qs bs (Const ("_constrain", _) $ t $ typ) =
 | 
| 42264 | 213 | (case Term_Position.decode_position typ of | 
| 42242 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 214 | SOME p => decode (p :: ps) qs bs t | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 215 | | NONE => Type.constraint (decodeT typ) (decode ps qs bs t)) | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 216 |           | decode ps qs bs (Const ("_constrainAbs", _) $ t $ typ) =
 | 
| 42264 | 217 | (case Term_Position.decode_position typ of | 
| 42242 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 218 | SOME q => decode ps (q :: qs) bs t | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 219 | | NONE => Type.constraint (decodeT typ --> dummyT) (decode ps qs bs t)) | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 220 | | decode _ qs bs (Abs (x, T, t)) = | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 221 | let | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 222 | val id = serial_string (); | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 223 | val _ = report qs (markup_bound true) id; | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 224 | in Abs (x, T, decode [] [] (id :: bs) t) end | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 225 | | decode _ _ bs (t $ u) = decode [] [] bs t $ decode [] [] bs u | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 226 | | decode ps _ _ (Const (a, T)) = | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 227 | (case try Lexicon.unmark_fixed a of | 
| 42282 
4fa41e068ff0
report literal tokens according to parsetree head;
 wenzelm parents: 
42281diff
changeset | 228 | 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 | 229 | | NONE => | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 230 | let | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 231 | val c = | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 232 | (case try Lexicon.unmark_const a of | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 233 | SOME c => c | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 234 | | NONE => snd (get_const a)); | 
| 42282 
4fa41e068ff0
report literal tokens according to parsetree head;
 wenzelm parents: 
42281diff
changeset | 235 | val _ = report ps (markup_const ctxt) c; | 
| 42242 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 236 | in Const (c, T) end) | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 237 | | decode ps _ _ (Free (a, T)) = | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 238 | (case (get_free a, get_const a) of | 
| 42282 
4fa41e068ff0
report literal tokens according to parsetree head;
 wenzelm parents: 
42281diff
changeset | 239 | (SOME x, _) => (report ps (markup_free ctxt) x; Free (x, T)) | 
| 
4fa41e068ff0
report literal tokens according to parsetree head;
 wenzelm parents: 
42281diff
changeset | 240 | | (_, (true, c)) => (report ps (markup_const ctxt) c; Const (c, T)) | 
| 42242 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 241 | | (_, (false, c)) => | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 242 | if Long_Name.is_qualified c | 
| 42282 
4fa41e068ff0
report literal tokens according to parsetree head;
 wenzelm parents: 
42281diff
changeset | 243 | then (report ps (markup_const ctxt) c; Const (c, T)) | 
| 
4fa41e068ff0
report literal tokens according to parsetree head;
 wenzelm parents: 
42281diff
changeset | 244 | else (report ps (markup_free ctxt) c; Free (c, T))) | 
| 42242 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 245 | | 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 | 246 | | decode ps _ bs (t as Bound i) = | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 247 | (case try (nth bs) i of | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 248 | SOME id => (report ps (markup_bound false) id; t) | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 249 | | NONE => t); | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 250 | |
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 251 | val tm' = Exn.interruptible_capture (fn () => decode [] [] [] tm) (); | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 252 | in (! reports, tm') end; | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 253 | |
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 254 | |
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 255 | |
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 256 | (** parse **) | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 257 | |
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 258 | (* results *) | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 259 | |
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 260 | fun ambiguity_msg pos = "Parse error: ambiguous syntax" ^ Position.str_of pos; | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 261 | |
| 43761 
e72ba84ae58f
tuned signature -- corresponding to Scala version;
 wenzelm parents: 
43731diff
changeset | 262 | 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 | 263 | 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 | 264 | |
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 265 | fun report_result ctxt pos results = | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 266 | (case (proper_results results, failed_results results) of | 
| 44736 | 267 | ([], (reports, exn) :: _) => (Context_Position.reports ctxt reports; reraise exn) | 
| 268 | | ([(reports, x)], _) => (Context_Position.reports ctxt reports; x) | |
| 42242 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 269 | | _ => error (ambiguity_msg pos)); | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 270 | |
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 271 | |
| 42249 | 272 | (* parse raw asts *) | 
| 42242 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 273 | |
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 274 | fun parse_asts ctxt raw root (syms, pos) = | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 275 | let | 
| 42360 | 276 | val syn = Proof_Context.syn_of ctxt; | 
| 42253 | 277 | val ast_tr = Syntax.parse_ast_translation syn; | 
| 42242 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 278 | |
| 42251 | 279 | val toks = Syntax.tokenize syn raw syms; | 
| 44736 | 280 | val _ = Context_Position.reports ctxt (map Lexicon.report_of_token toks); | 
| 42242 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 281 | |
| 42251 | 282 | val pts = Syntax.parse ctxt syn root (filter Lexicon.is_proper toks) | 
| 42242 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 283 | handle ERROR msg => | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 284 | error (msg ^ | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 285 | implode (map (Markup.markup Markup.report o Lexicon.reported_token_range ctxt) toks)); | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 286 | val len = length pts; | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 287 | |
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 288 | val limit = Config.get ctxt Syntax.ambiguity_limit; | 
| 44069 | 289 | val warnings = Config.get ctxt Syntax.ambiguity_warnings; | 
| 42242 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 290 | val _ = | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 291 | if len <= Config.get ctxt Syntax.ambiguity_level then () | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 292 | else if not (Config.get ctxt Syntax.ambiguity_enabled) then error (ambiguity_msg pos) | 
| 44069 | 293 | else if warnings then | 
| 42242 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 294 | (Context_Position.if_visible ctxt warning (cat_lines | 
| 44564 
96ba83710946
tuned positions of ambiguity messages -- less intrusive in IDE view;
 wenzelm parents: 
44069diff
changeset | 295 |           (("Ambiguous input" ^ Position.str_of (Position.reset_range pos) ^
 | 
| 42242 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 296 | "\nproduces " ^ string_of_int len ^ " parse trees" ^ | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 297 |             (if len <= limit then "" else " (" ^ string_of_int limit ^ " displayed)") ^ ":") ::
 | 
| 44069 | 298 | map (Pretty.string_of o Parser.pretty_parsetree) (take limit pts)))) | 
| 299 | else (); | |
| 42242 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 300 | |
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 301 | val constrain_pos = not raw andalso Config.get ctxt Syntax.positions; | 
| 42253 | 302 | val parsetree_to_ast = parsetree_to_ast ctxt constrain_pos ast_tr; | 
| 42242 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 303 | in map parsetree_to_ast pts end; | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 304 | |
| 42249 | 305 | fun parse_raw ctxt root input = | 
| 42242 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 306 | let | 
| 42360 | 307 | val syn = Proof_Context.syn_of ctxt; | 
| 42253 | 308 | val tr = Syntax.parse_translation syn; | 
| 42255 | 309 | val parse_rules = Syntax.parse_rules syn; | 
| 42242 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 310 | in | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 311 | parse_asts ctxt false root input | 
| 42255 | 312 | |> (map o apsnd o Exn.maps_result) | 
| 313 | (Ast.normalize ctxt parse_rules #> Exn.interruptible_capture (ast_to_term ctxt tr)) | |
| 42242 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 314 | end; | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 315 | |
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 316 | |
| 42249 | 317 | (* parse logical entities *) | 
| 42242 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 318 | |
| 42241 
dd8029f71e1c
separate module for standard implementation of inner syntax operations;
 wenzelm parents: diff
changeset | 319 | fun parse_failed ctxt pos msg kind = | 
| 
dd8029f71e1c
separate module for standard implementation of inner syntax operations;
 wenzelm parents: diff
changeset | 320 |   cat_error msg ("Failed to parse " ^ kind ^
 | 
| 
dd8029f71e1c
separate module for standard implementation of inner syntax operations;
 wenzelm parents: diff
changeset | 321 | Markup.markup Markup.report (Context_Position.reported_text ctxt pos Markup.bad "")); | 
| 
dd8029f71e1c
separate module for standard implementation of inner syntax operations;
 wenzelm parents: diff
changeset | 322 | |
| 43731 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43552diff
changeset | 323 | fun parse_sort ctxt = | 
| 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43552diff
changeset | 324 | Syntax.parse_token ctxt Term_XML.Decode.sort Markup.sort | 
| 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43552diff
changeset | 325 | (fn (syms, pos) => | 
| 42249 | 326 | parse_raw ctxt "sort" (syms, pos) | 
| 327 | |> report_result ctxt pos | |
| 328 | |> sort_of_term | |
| 43731 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43552diff
changeset | 329 | |> 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 | 330 | handle ERROR msg => parse_failed ctxt pos msg "sort"); | 
| 42241 
dd8029f71e1c
separate module for standard implementation of inner syntax operations;
 wenzelm parents: diff
changeset | 331 | |
| 43731 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43552diff
changeset | 332 | fun parse_typ ctxt = | 
| 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43552diff
changeset | 333 | Syntax.parse_token ctxt Term_XML.Decode.typ Markup.typ | 
| 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43552diff
changeset | 334 | (fn (syms, pos) => | 
| 42249 | 335 | parse_raw ctxt "type" (syms, pos) | 
| 336 | |> report_result ctxt pos | |
| 42360 | 337 | |> (fn t => typ_of_term (Proof_Context.get_sort ctxt (term_sorts t)) t) | 
| 43731 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43552diff
changeset | 338 | handle ERROR msg => parse_failed ctxt pos msg "type"); | 
| 42241 
dd8029f71e1c
separate module for standard implementation of inner syntax operations;
 wenzelm parents: diff
changeset | 339 | |
| 43731 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43552diff
changeset | 340 | fun parse_term is_prop ctxt = | 
| 42241 
dd8029f71e1c
separate module for standard implementation of inner syntax operations;
 wenzelm parents: diff
changeset | 341 | let | 
| 42281 | 342 | val (markup, kind, root, constrain) = | 
| 343 | if is_prop | |
| 344 | then (Markup.prop, "proposition", "prop", Type.constraint propT) | |
| 42288 
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
 wenzelm parents: 
42284diff
changeset | 345 | else (Markup.term, "term", Config.get ctxt Syntax.root, I); | 
| 43731 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43552diff
changeset | 346 | val decode = constrain o Term_XML.Decode.term; | 
| 42249 | 347 | in | 
| 43731 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43552diff
changeset | 348 | Syntax.parse_token ctxt decode markup | 
| 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43552diff
changeset | 349 | (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 | 350 | let | 
| 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43552diff
changeset | 351 | val results = parse_raw ctxt root (syms, pos) |> map (decode_term ctxt); | 
| 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43552diff
changeset | 352 | val ambiguity = length (proper_results results); | 
| 42249 | 353 | |
| 43731 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43552diff
changeset | 354 | val level = Config.get ctxt Syntax.ambiguity_level; | 
| 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43552diff
changeset | 355 | val limit = Config.get ctxt Syntax.ambiguity_limit; | 
| 44069 | 356 | val warnings = Config.get ctxt Syntax.ambiguity_warnings; | 
| 42249 | 357 | |
| 43731 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43552diff
changeset | 358 | val ambig_msg = | 
| 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43552diff
changeset | 359 | if ambiguity > 1 andalso ambiguity <= level then | 
| 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43552diff
changeset | 360 | ["Got more than one parse tree.\n\ | 
| 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43552diff
changeset | 361 | \Retry with smaller syntax_ambiguity_level for more information."] | 
| 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43552diff
changeset | 362 | else []; | 
| 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43552diff
changeset | 363 | |
| 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43552diff
changeset | 364 | (*brute-force disambiguation via type-inference*) | 
| 43761 
e72ba84ae58f
tuned signature -- corresponding to Scala version;
 wenzelm parents: 
43731diff
changeset | 365 | fun check t = (Syntax.check_term ctxt (constrain t); Exn.Res t) | 
| 43731 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43552diff
changeset | 366 | handle exn as ERROR _ => Exn.Exn exn; | 
| 42249 | 367 | |
| 43731 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43552diff
changeset | 368 | val results' = | 
| 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43552diff
changeset | 369 | if ambiguity > 1 then | 
| 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43552diff
changeset | 370 | (Par_List.map_name "Syntax_Phases.parse_term" o apsnd o Exn.maps_result) | 
| 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43552diff
changeset | 371 | check results | 
| 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43552diff
changeset | 372 | else results; | 
| 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43552diff
changeset | 373 | val reports' = fst (hd results'); | 
| 42249 | 374 | |
| 43731 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43552diff
changeset | 375 | 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 | 376 | val checked = map snd (proper_results results'); | 
| 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43552diff
changeset | 377 | val len = length checked; | 
| 42249 | 378 | |
| 43731 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43552diff
changeset | 379 | val show_term = Syntax.string_of_term (Config.put Printer.show_brackets true ctxt); | 
| 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43552diff
changeset | 380 | in | 
| 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43552diff
changeset | 381 | if len = 0 then | 
| 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43552diff
changeset | 382 | report_result ctxt pos | 
| 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43552diff
changeset | 383 | [(reports', Exn.Exn (Exn.EXCEPTIONS (map ERROR ambig_msg @ errs)))] | 
| 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43552diff
changeset | 384 | else if len = 1 then | 
| 44069 | 385 | (if ambiguity > level andalso warnings then | 
| 43731 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43552diff
changeset | 386 | Context_Position.if_visible ctxt warning | 
| 44564 
96ba83710946
tuned positions of ambiguity messages -- less intrusive in IDE view;
 wenzelm parents: 
44069diff
changeset | 387 |                 ("Fortunately, only one parse tree is type correct" ^
 | 
| 
96ba83710946
tuned positions of ambiguity messages -- less intrusive in IDE view;
 wenzelm parents: 
44069diff
changeset | 388 | Position.str_of (Position.reset_range pos) ^ | 
| 
96ba83710946
tuned positions of ambiguity messages -- less intrusive in IDE view;
 wenzelm parents: 
44069diff
changeset | 389 | ",\nbut you may still want to disambiguate your grammar or your input.") | 
| 43731 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43552diff
changeset | 390 | else (); report_result ctxt pos results') | 
| 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43552diff
changeset | 391 | else | 
| 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43552diff
changeset | 392 | report_result ctxt pos | 
| 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43552diff
changeset | 393 | [(reports', Exn.Exn (ERROR (cat_lines (ambig_msg @ | 
| 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43552diff
changeset | 394 |                 (("Ambiguous input, " ^ string_of_int len ^ " terms are type correct" ^
 | 
| 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43552diff
changeset | 395 |                   (if len <= limit then "" else " (" ^ string_of_int limit ^ " displayed)") ^ ":") ::
 | 
| 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43552diff
changeset | 396 | map show_term (take limit checked))))))] | 
| 
70072780e095
inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
 wenzelm parents: 
43552diff
changeset | 397 | end handle ERROR msg => parse_failed ctxt pos msg kind) | 
| 42249 | 398 | end; | 
| 42241 
dd8029f71e1c
separate module for standard implementation of inner syntax operations;
 wenzelm parents: diff
changeset | 399 | |
| 
dd8029f71e1c
separate module for standard implementation of inner syntax operations;
 wenzelm parents: diff
changeset | 400 | |
| 42242 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 401 | (* parse_ast_pattern *) | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 402 | |
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 403 | fun parse_ast_pattern ctxt (root, str) = | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 404 | let | 
| 42360 | 405 | val syn = Proof_Context.syn_of ctxt; | 
| 42242 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 406 | |
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 407 | fun constify (ast as Ast.Constant _) = ast | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 408 | | constify (ast as 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 | 409 | if is_some (Syntax.lookup_const syn x) orelse Long_Name.is_qualified x | 
| 
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
 wenzelm parents: 
42296diff
changeset | 410 | then Ast.Constant x | 
| 42242 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 411 | else ast | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 412 | | constify (Ast.Appl asts) = Ast.Appl (map constify asts); | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 413 | |
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 414 | val (syms, pos) = Syntax.read_token str; | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 415 | in | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 416 | parse_asts ctxt true root (syms, pos) | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 417 | |> report_result ctxt pos | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 418 | |> constify | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 419 | end; | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 420 | |
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 421 | |
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 422 | |
| 42245 | 423 | (** encode parse trees **) | 
| 424 | ||
| 425 | (* term_of_sort *) | |
| 426 | ||
| 427 | fun term_of_sort S = | |
| 428 | let | |
| 42476 | 429 | val class = Syntax.const o Lexicon.mark_class; | 
| 42245 | 430 | |
| 431 | fun classes [c] = class c | |
| 42476 | 432 | | classes (c :: cs) = Syntax.const "_classes" $ class c $ classes cs; | 
| 42245 | 433 | in | 
| 434 | (case S of | |
| 42476 | 435 | [] => Syntax.const "_topsort" | 
| 42245 | 436 | | [c] => class c | 
| 42476 | 437 | | cs => Syntax.const "_sort" $ classes cs) | 
| 42245 | 438 | end; | 
| 439 | ||
| 440 | ||
| 441 | (* term_of_typ *) | |
| 442 | ||
| 42247 
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
 wenzelm parents: 
42245diff
changeset | 443 | fun term_of_typ ctxt ty = | 
| 42245 | 444 | let | 
| 42247 
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
 wenzelm parents: 
42245diff
changeset | 445 | val show_sorts = Config.get ctxt show_sorts; | 
| 
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
 wenzelm parents: 
42245diff
changeset | 446 | |
| 42245 | 447 | fun of_sort t S = | 
| 42476 | 448 | if show_sorts then Syntax.const "_ofsort" $ t $ term_of_sort S | 
| 42245 | 449 | else t; | 
| 450 | ||
| 451 | fun term_of (Type (a, Ts)) = | |
| 42476 | 452 | Term.list_comb (Syntax.const (Lexicon.mark_type a), map term_of Ts) | 
| 42245 | 453 | | term_of (TFree (x, S)) = | 
| 42290 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42289diff
changeset | 454 | if is_some (Term_Position.decode x) then Syntax.free x | 
| 42476 | 455 | else of_sort (Syntax.const "_tfree" $ Syntax.free x) S | 
| 456 | | term_of (TVar (xi, S)) = of_sort (Syntax.const "_tvar" $ Syntax.var xi) S; | |
| 42245 | 457 | in term_of ty end; | 
| 458 | ||
| 459 | ||
| 460 | (* simple_ast_of *) | |
| 461 | ||
| 462 | fun simple_ast_of ctxt = | |
| 463 | let | |
| 464 | val tune_var = if Config.get ctxt show_question_marks then I else unprefix "?"; | |
| 465 | fun ast_of (Const (c, _)) = Ast.Constant c | |
| 466 | | ast_of (Free (x, _)) = Ast.Variable x | |
| 467 | | ast_of (Var (xi, _)) = Ast.Variable (tune_var (Term.string_of_vname xi)) | |
| 468 | | ast_of (t as _ $ _) = | |
| 469 | let val (f, args) = strip_comb t | |
| 470 | in Ast.mk_appl (ast_of f) (map ast_of args) end | |
| 42408 | 471 |       | ast_of (Bound i) = Ast.Appl [Ast.Constant "_loose", Ast.Variable ("B." ^ string_of_int i)]
 | 
| 42245 | 472 | | ast_of (Abs _) = raise Fail "simple_ast_of: Abs"; | 
| 473 | in ast_of end; | |
| 474 | ||
| 475 | ||
| 476 | (* sort_to_ast and typ_to_ast *) | |
| 477 | ||
| 478 | fun ast_of_termT ctxt trf tm = | |
| 479 | let | |
| 480 | val ctxt' = Config.put show_sorts false ctxt; | |
| 481 |     fun ast_of (t as Const ("_tfree", _) $ Free _) = simple_ast_of ctxt t
 | |
| 482 |       | ast_of (t as Const ("_tvar", _) $ Var _) = simple_ast_of ctxt t
 | |
| 483 | | ast_of (Const (a, _)) = trans a [] | |
| 484 | | ast_of (t as _ $ _) = | |
| 485 | (case strip_comb t of | |
| 486 | (Const (a, _), args) => trans a args | |
| 487 | | (f, args) => Ast.Appl (map ast_of (f :: args))) | |
| 488 | | ast_of t = simple_ast_of ctxt t | |
| 42254 | 489 | and trans a args = ast_of (trf a ctxt' dummyT args) | 
| 490 | handle Match => Ast.mk_appl (Ast.Constant a) (map ast_of args); | |
| 42245 | 491 | in ast_of tm end; | 
| 492 | ||
| 493 | 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 | 494 | fun typ_to_ast ctxt trf T = ast_of_termT ctxt trf (term_of_typ ctxt T); | 
| 42245 | 495 | |
| 496 | ||
| 497 | (* term_to_ast *) | |
| 498 | ||
| 42252 | 499 | fun term_to_ast idents is_syntax_const ctxt trf tm = | 
| 42245 | 500 | let | 
| 501 | val show_types = | |
| 502 | Config.get ctxt show_types orelse Config.get ctxt show_sorts orelse | |
| 503 | Config.get ctxt show_all_types; | |
| 504 | val show_structs = Config.get ctxt show_structs; | |
| 505 | val show_free_types = Config.get ctxt show_free_types; | |
| 506 | val show_all_types = Config.get ctxt show_all_types; | |
| 507 | ||
| 508 |     val {structs, fixes} = idents;
 | |
| 509 | ||
| 42249 | 510 | fun mark_atoms ((t as Const (c, _)) $ u) = | 
| 42294 | 511 | if member (op =) Syntax.token_markers c | 
| 42245 | 512 | then t $ u else mark_atoms t $ mark_atoms u | 
| 513 | | mark_atoms (t $ u) = mark_atoms t $ mark_atoms u | |
| 514 | | mark_atoms (Abs (x, T, t)) = Abs (x, T, mark_atoms t) | |
| 515 | | mark_atoms (t as Const (c, T)) = | |
| 42252 | 516 | if is_syntax_const c then t | 
| 42245 | 517 | else Const (Lexicon.mark_const c, T) | 
| 518 | | mark_atoms (t as Free (x, T)) = | |
| 519 | let val i = find_index (fn s => s = x) structs + 1 in | |
| 520 | if i = 0 andalso member (op =) fixes x then | |
| 521 | Const (Lexicon.mark_fixed x, T) | |
| 522 | else if i = 1 andalso not show_structs then | |
| 42476 | 523 | Syntax.const "_struct" $ Syntax.const "_indexdefault" | 
| 524 | else Syntax.const "_free" $ t | |
| 42245 | 525 | end | 
| 526 | | mark_atoms (t as Var (xi, T)) = | |
| 42288 
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
 wenzelm parents: 
42284diff
changeset | 527 |           if xi = Syntax_Ext.dddot_indexname then Const ("_DDDOT", T)
 | 
| 42476 | 528 | else Syntax.const "_var" $ t | 
| 42245 | 529 | | mark_atoms a = a; | 
| 530 | ||
| 531 | fun prune_typs (t_seen as (Const _, _)) = t_seen | |
| 532 | | prune_typs (t as Free (x, ty), seen) = | |
| 533 | if ty = dummyT then (t, seen) | |
| 42290 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42289diff
changeset | 534 | else if not show_free_types orelse member (op aconv) seen t then (Syntax.free x, seen) | 
| 42245 | 535 | else (t, t :: seen) | 
| 536 | | prune_typs (t as Var (xi, ty), seen) = | |
| 537 | if ty = dummyT then (t, seen) | |
| 42290 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42289diff
changeset | 538 | else if not show_free_types orelse member (op aconv) seen t then (Syntax.var xi, seen) | 
| 42245 | 539 | else (t, t :: seen) | 
| 540 | | prune_typs (t_seen as (Bound _, _)) = t_seen | |
| 541 | | prune_typs (Abs (x, ty, t), seen) = | |
| 542 | let val (t', seen') = prune_typs (t, seen); | |
| 543 | in (Abs (x, ty, t'), seen') end | |
| 544 | | prune_typs (t1 $ t2, seen) = | |
| 545 | let | |
| 546 | val (t1', seen') = prune_typs (t1, seen); | |
| 547 | val (t2', seen'') = prune_typs (t2, seen'); | |
| 548 | in (t1' $ t2', seen'') end; | |
| 549 | ||
| 550 | fun ast_of tm = | |
| 551 | (case strip_comb tm of | |
| 42284 | 552 | (t as Abs _, ts) => Ast.mk_appl (ast_of (Syntax_Trans.abs_tr' ctxt t)) (map ast_of ts) | 
| 42245 | 553 |       | ((c as Const ("_free", _)), Free (x, T) :: ts) =>
 | 
| 42290 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42289diff
changeset | 554 | Ast.mk_appl (constrain (c $ Syntax.free x) T) (map ast_of ts) | 
| 42245 | 555 |       | ((c as Const ("_var", _)), Var (xi, T) :: ts) =>
 | 
| 42290 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42289diff
changeset | 556 | Ast.mk_appl (constrain (c $ Syntax.var xi) T) (map ast_of ts) | 
| 42245 | 557 |       | ((c as Const ("_bound", _)), Free (x, T) :: ts) =>
 | 
| 42290 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42289diff
changeset | 558 | Ast.mk_appl (constrain (c $ Syntax.free x) T) (map ast_of ts) | 
| 42245 | 559 |       | (Const ("_idtdummy", T), ts) =>
 | 
| 42476 | 560 | Ast.mk_appl (constrain (Syntax.const "_idtdummy") T) (map ast_of ts) | 
| 42245 | 561 | | (const as Const (c, T), ts) => | 
| 562 | if show_all_types | |
| 563 | then Ast.mk_appl (constrain const T) (map ast_of ts) | |
| 564 | else trans c T ts | |
| 565 | | (t, ts) => Ast.mk_appl (simple_ast_of ctxt t) (map ast_of ts)) | |
| 566 | ||
| 42254 | 567 | and trans a T args = ast_of (trf a ctxt T args) | 
| 568 | handle Match => Ast.mk_appl (Ast.Constant a) (map ast_of args) | |
| 42245 | 569 | |
| 570 | and constrain t T = | |
| 571 | if show_types andalso T <> dummyT then | |
| 42248 | 572 | Ast.Appl [Ast.Constant "_constrain", simple_ast_of ctxt t, | 
| 42247 
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
 wenzelm parents: 
42245diff
changeset | 573 | ast_of_termT ctxt trf (term_of_typ ctxt T)] | 
| 42245 | 574 | else simple_ast_of ctxt t; | 
| 575 | in | |
| 576 | tm | |
| 42284 | 577 | |> Syntax_Trans.prop_tr' | 
| 42245 | 578 | |> show_types ? (#1 o prune_typs o rpair []) | 
| 579 | |> mark_atoms | |
| 580 | |> ast_of | |
| 581 | end; | |
| 582 | ||
| 583 | ||
| 584 | ||
| 42242 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 585 | (** unparse **) | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 586 | |
| 42245 | 587 | local | 
| 588 | ||
| 42267 
9566078ad905
simplified printer context: uniform externing and static token translations;
 wenzelm parents: 
42264diff
changeset | 589 | fun free_or_skolem ctxt x = | 
| 
9566078ad905
simplified printer context: uniform externing and static token translations;
 wenzelm parents: 
42264diff
changeset | 590 | let | 
| 
9566078ad905
simplified printer context: uniform externing and static token translations;
 wenzelm parents: 
42264diff
changeset | 591 | val m = | 
| 
9566078ad905
simplified printer context: uniform externing and static token translations;
 wenzelm parents: 
42264diff
changeset | 592 | if Variable.is_fixed ctxt x orelse Syntax.is_pretty_global ctxt | 
| 
9566078ad905
simplified printer context: uniform externing and static token translations;
 wenzelm parents: 
42264diff
changeset | 593 | then Markup.fixed x | 
| 
9566078ad905
simplified printer context: uniform externing and static token translations;
 wenzelm parents: 
42264diff
changeset | 594 | else Markup.hilite; | 
| 
9566078ad905
simplified printer context: uniform externing and static token translations;
 wenzelm parents: 
42264diff
changeset | 595 | in | 
| 
9566078ad905
simplified printer context: uniform externing and static token translations;
 wenzelm parents: 
42264diff
changeset | 596 | if can Name.dest_skolem x | 
| 42488 
4638622bcaa1
reorganized fixes as specialized (global) name space;
 wenzelm parents: 
42476diff
changeset | 597 | then ([m, Markup.skolem], Variable.revert_fixed ctxt x) | 
| 42267 
9566078ad905
simplified printer context: uniform externing and static token translations;
 wenzelm parents: 
42264diff
changeset | 598 | else ([m, Markup.free], x) | 
| 
9566078ad905
simplified printer context: uniform externing and static token translations;
 wenzelm parents: 
42264diff
changeset | 599 | end; | 
| 
9566078ad905
simplified printer context: uniform externing and static token translations;
 wenzelm parents: 
42264diff
changeset | 600 | |
| 
9566078ad905
simplified printer context: uniform externing and static token translations;
 wenzelm parents: 
42264diff
changeset | 601 | fun var_or_skolem s = | 
| 
9566078ad905
simplified printer context: uniform externing and static token translations;
 wenzelm parents: 
42264diff
changeset | 602 | (case Lexicon.read_variable s of | 
| 
9566078ad905
simplified printer context: uniform externing and static token translations;
 wenzelm parents: 
42264diff
changeset | 603 | SOME (x, i) => | 
| 
9566078ad905
simplified printer context: uniform externing and static token translations;
 wenzelm parents: 
42264diff
changeset | 604 | (case try Name.dest_skolem x of | 
| 
9566078ad905
simplified printer context: uniform externing and static token translations;
 wenzelm parents: 
42264diff
changeset | 605 | NONE => (Markup.var, s) | 
| 
9566078ad905
simplified printer context: uniform externing and static token translations;
 wenzelm parents: 
42264diff
changeset | 606 | | SOME x' => (Markup.skolem, Term.string_of_vname (x', i))) | 
| 
9566078ad905
simplified printer context: uniform externing and static token translations;
 wenzelm parents: 
42264diff
changeset | 607 | | NONE => (Markup.var, s)); | 
| 
9566078ad905
simplified printer context: uniform externing and static token translations;
 wenzelm parents: 
42264diff
changeset | 608 | |
| 
9566078ad905
simplified printer context: uniform externing and static token translations;
 wenzelm parents: 
42264diff
changeset | 609 | fun unparse_t t_to_ast prt_t markup ctxt t = | 
| 42245 | 610 | let | 
| 42360 | 611 | val syn = Proof_Context.syn_of ctxt; | 
| 42267 
9566078ad905
simplified printer context: uniform externing and static token translations;
 wenzelm parents: 
42264diff
changeset | 612 | |
| 
9566078ad905
simplified printer context: uniform externing and static token translations;
 wenzelm parents: 
42264diff
changeset | 613 | fun token_trans "_tfree" x = SOME (Pretty.mark_str (Markup.tfree, x)) | 
| 
9566078ad905
simplified printer context: uniform externing and static token translations;
 wenzelm parents: 
42264diff
changeset | 614 | | token_trans "_tvar" x = SOME (Pretty.mark_str (Markup.tvar, x)) | 
| 
9566078ad905
simplified printer context: uniform externing and static token translations;
 wenzelm parents: 
42264diff
changeset | 615 | | token_trans "_free" x = SOME (Pretty.marks_str (free_or_skolem ctxt x)) | 
| 
9566078ad905
simplified printer context: uniform externing and static token translations;
 wenzelm parents: 
42264diff
changeset | 616 | | token_trans "_bound" x = SOME (Pretty.mark_str (Markup.bound, x)) | 
| 42408 | 617 | | token_trans "_loose" x = SOME (Pretty.mark_str (Markup.malformed, x)) | 
| 42267 
9566078ad905
simplified printer context: uniform externing and static token translations;
 wenzelm parents: 
42264diff
changeset | 618 | | token_trans "_var" x = SOME (Pretty.mark_str (var_or_skolem x)) | 
| 
9566078ad905
simplified printer context: uniform externing and static token translations;
 wenzelm parents: 
42264diff
changeset | 619 | | token_trans "_numeral" x = SOME (Pretty.mark_str (Markup.numeral, x)) | 
| 
9566078ad905
simplified printer context: uniform externing and static token translations;
 wenzelm parents: 
42264diff
changeset | 620 | | token_trans "_inner_string" x = SOME (Pretty.mark_str (Markup.inner_string, x)) | 
| 
9566078ad905
simplified printer context: uniform externing and static token translations;
 wenzelm parents: 
42264diff
changeset | 621 | | token_trans _ _ = NONE; | 
| 
9566078ad905
simplified printer context: uniform externing and static token translations;
 wenzelm parents: 
42264diff
changeset | 622 | |
| 42301 
d7ca0c03d4ea
unparse: more accurate markup for syntax consts, notably binders;
 wenzelm parents: 
42298diff
changeset | 623 | fun markup_extern c = | 
| 
d7ca0c03d4ea
unparse: more accurate markup for syntax consts, notably binders;
 wenzelm parents: 
42298diff
changeset | 624 | (case Syntax.lookup_const syn c of | 
| 
d7ca0c03d4ea
unparse: more accurate markup for syntax consts, notably binders;
 wenzelm parents: 
42298diff
changeset | 625 | SOME "" => ([], c) | 
| 
d7ca0c03d4ea
unparse: more accurate markup for syntax consts, notably binders;
 wenzelm parents: 
42298diff
changeset | 626 | | SOME b => markup_extern b | 
| 
d7ca0c03d4ea
unparse: more accurate markup for syntax consts, notably binders;
 wenzelm parents: 
42298diff
changeset | 627 | | NONE => c |> Lexicon.unmark | 
| 43548 
f231a7594e54
type classes: entity markup instead of old-style token markup;
 wenzelm parents: 
42493diff
changeset | 628 |          {case_class = fn x => (markup_class ctxt x, Proof_Context.extern_class ctxt x),
 | 
| 43552 | 629 | case_type = fn x => (markup_type ctxt x, Proof_Context.extern_type ctxt x), | 
| 630 | case_const = fn x => (markup_const ctxt x, Proof_Context.extern_const ctxt x), | |
| 42301 
d7ca0c03d4ea
unparse: more accurate markup for syntax consts, notably binders;
 wenzelm parents: 
42298diff
changeset | 631 | case_fixed = fn x => free_or_skolem ctxt x, | 
| 
d7ca0c03d4ea
unparse: more accurate markup for syntax consts, notably binders;
 wenzelm parents: 
42298diff
changeset | 632 | case_default = fn x => ([], x)}); | 
| 42245 | 633 | in | 
| 42255 | 634 | t_to_ast ctxt (Syntax.print_translation syn) t | 
| 635 | |> Ast.normalize ctxt (Syntax.print_rules syn) | |
| 42267 
9566078ad905
simplified printer context: uniform externing and static token translations;
 wenzelm parents: 
42264diff
changeset | 636 | |> prt_t ctxt (Syntax.prtabs syn) (Syntax.print_ast_translation syn) token_trans markup_extern | 
| 42253 | 637 | |> Pretty.markup markup | 
| 42245 | 638 | end; | 
| 639 | ||
| 640 | in | |
| 641 | ||
| 42267 
9566078ad905
simplified printer context: uniform externing and static token translations;
 wenzelm parents: 
42264diff
changeset | 642 | val unparse_sort = unparse_t sort_to_ast Printer.pretty_typ_ast Markup.sort; | 
| 
9566078ad905
simplified printer context: uniform externing and static token translations;
 wenzelm parents: 
42264diff
changeset | 643 | val unparse_typ = unparse_t typ_to_ast Printer.pretty_typ_ast Markup.typ; | 
| 42241 
dd8029f71e1c
separate module for standard implementation of inner syntax operations;
 wenzelm parents: diff
changeset | 644 | |
| 
dd8029f71e1c
separate module for standard implementation of inner syntax operations;
 wenzelm parents: diff
changeset | 645 | fun unparse_term ctxt = | 
| 
dd8029f71e1c
separate module for standard implementation of inner syntax operations;
 wenzelm parents: diff
changeset | 646 | let | 
| 42360 | 647 | val thy = Proof_Context.theory_of ctxt; | 
| 648 | val syn = Proof_Context.syn_of ctxt; | |
| 649 | val idents = Local_Syntax.idents_of (Proof_Context.syntax_of ctxt); | |
| 42241 
dd8029f71e1c
separate module for standard implementation of inner syntax operations;
 wenzelm parents: diff
changeset | 650 | in | 
| 42298 
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
 wenzelm parents: 
42296diff
changeset | 651 | unparse_t (term_to_ast idents (is_some o Syntax.lookup_const syn)) | 
| 42267 
9566078ad905
simplified printer context: uniform externing and static token translations;
 wenzelm parents: 
42264diff
changeset | 652 | (Printer.pretty_term_ast (not (Pure_Thy.old_appl_syntax thy))) | 
| 
9566078ad905
simplified printer context: uniform externing and static token translations;
 wenzelm parents: 
42264diff
changeset | 653 | Markup.term ctxt | 
| 42241 
dd8029f71e1c
separate module for standard implementation of inner syntax operations;
 wenzelm parents: diff
changeset | 654 | end; | 
| 
dd8029f71e1c
separate module for standard implementation of inner syntax operations;
 wenzelm parents: diff
changeset | 655 | |
| 42249 | 656 | end; | 
| 657 | ||
| 42241 
dd8029f71e1c
separate module for standard implementation of inner syntax operations;
 wenzelm parents: diff
changeset | 658 | |
| 42245 | 659 | |
| 660 | (** translations **) | |
| 661 | ||
| 662 | (* type propositions *) | |
| 663 | ||
| 42247 
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
 wenzelm parents: 
42245diff
changeset | 664 | fun type_prop_tr' ctxt T [Const ("\\<^const>Pure.sort_constraint", _)] =
 | 
| 42476 | 665 | 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 | 666 | | type_prop_tr' ctxt T [t] = | 
| 42476 | 667 | Syntax.const "_ofclass" $ term_of_typ ctxt T $ t | 
| 42245 | 668 |   | type_prop_tr' _ T ts = raise TYPE ("type_prop_tr'", [T], ts);
 | 
| 669 | ||
| 670 | ||
| 671 | (* type reflection *) | |
| 672 | ||
| 42247 
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
 wenzelm parents: 
42245diff
changeset | 673 | fun type_tr' ctxt (Type ("itself", [T])) ts =
 | 
| 42476 | 674 | Term.list_comb (Syntax.const "_TYPE" $ term_of_typ ctxt T, ts) | 
| 42245 | 675 | | type_tr' _ _ _ = raise Match; | 
| 676 | ||
| 677 | ||
| 678 | (* type constraints *) | |
| 679 | ||
| 42247 
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
 wenzelm parents: 
42245diff
changeset | 680 | fun type_constraint_tr' ctxt (Type ("fun", [T, _])) (t :: ts) =
 | 
| 42476 | 681 | Term.list_comb (Syntax.const "_constrain" $ t $ term_of_typ ctxt T, ts) | 
| 42245 | 682 | | type_constraint_tr' _ _ _ = raise Match; | 
| 683 | ||
| 684 | ||
| 42295 
8fdbb3b10beb
moved CONST syntax/translations to their proper place;
 wenzelm parents: 
42294diff
changeset | 685 | (* authentic syntax *) | 
| 
8fdbb3b10beb
moved CONST syntax/translations to their proper place;
 wenzelm parents: 
42294diff
changeset | 686 | |
| 
8fdbb3b10beb
moved CONST syntax/translations to their proper place;
 wenzelm parents: 
42294diff
changeset | 687 | fun const_ast_tr intern ctxt [Ast.Variable c] = | 
| 
8fdbb3b10beb
moved CONST syntax/translations to their proper place;
 wenzelm parents: 
42294diff
changeset | 688 | let | 
| 42360 | 689 | val Const (c', _) = Proof_Context.read_const_proper ctxt false c; | 
| 42295 
8fdbb3b10beb
moved CONST syntax/translations to their proper place;
 wenzelm parents: 
42294diff
changeset | 690 | val d = if intern then Lexicon.mark_const c' else c; | 
| 
8fdbb3b10beb
moved CONST syntax/translations to their proper place;
 wenzelm parents: 
42294diff
changeset | 691 | in Ast.Constant d end | 
| 42470 | 692 | | const_ast_tr intern ctxt [Ast.Appl [Ast.Constant "_constrain", x, T as Ast.Variable pos]] = | 
| 693 | (Ast.Appl [Ast.Constant "_constrain", const_ast_tr intern ctxt [x], T] | |
| 694 | handle ERROR msg => | |
| 695 | error (msg ^ Position.str_of (the_default Position.none (Term_Position.decode pos)))) | |
| 42295 
8fdbb3b10beb
moved CONST syntax/translations to their proper place;
 wenzelm parents: 
42294diff
changeset | 696 |   | const_ast_tr _ _ asts = raise Ast.AST ("const_ast_tr", asts);
 | 
| 
8fdbb3b10beb
moved CONST syntax/translations to their proper place;
 wenzelm parents: 
42294diff
changeset | 697 | |
| 
8fdbb3b10beb
moved CONST syntax/translations to their proper place;
 wenzelm parents: 
42294diff
changeset | 698 | |
| 42245 | 699 | (* setup translations *) | 
| 700 | ||
| 701 | val _ = Context.>> (Context.map_theory | |
| 42295 
8fdbb3b10beb
moved CONST syntax/translations to their proper place;
 wenzelm parents: 
42294diff
changeset | 702 | (Sign.add_advanced_trfuns | 
| 
8fdbb3b10beb
moved CONST syntax/translations to their proper place;
 wenzelm parents: 
42294diff
changeset | 703 |   ([("_context_const", const_ast_tr true),
 | 
| 
8fdbb3b10beb
moved CONST syntax/translations to their proper place;
 wenzelm parents: 
42294diff
changeset | 704 |     ("_context_xconst", const_ast_tr false)], [], [], []) #>
 | 
| 
8fdbb3b10beb
moved CONST syntax/translations to their proper place;
 wenzelm parents: 
42294diff
changeset | 705 | Sign.add_advanced_trfunsT | 
| 42245 | 706 |    [("_type_prop", type_prop_tr'),
 | 
| 707 |     ("\\<^const>TYPE", type_tr'),
 | |
| 708 |     ("_type_constraint_", type_constraint_tr')]));
 | |
| 709 | ||
| 710 | ||
| 711 | ||
| 42382 
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
 wenzelm parents: 
42379diff
changeset | 712 | (** check/uncheck **) | 
| 
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
 wenzelm parents: 
42379diff
changeset | 713 | |
| 43794 
49cbbe2768a8
sub-structural sharing after Syntax.check phase, with global interning of logical entities (the latter is relevant when bypassing default parsing via YXML);
 wenzelm parents: 
43761diff
changeset | 714 | fun check_typs ctxt = Syntax.typ_check ctxt #> Term_Sharing.typs (Proof_Context.theory_of ctxt); | 
| 
49cbbe2768a8
sub-structural sharing after Syntax.check phase, with global interning of logical entities (the latter is relevant when bypassing default parsing via YXML);
 wenzelm parents: 
43761diff
changeset | 715 | fun check_terms ctxt = Syntax.term_check ctxt #> Term_Sharing.terms (Proof_Context.theory_of ctxt); | 
| 42382 
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
 wenzelm parents: 
42379diff
changeset | 716 | 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 | 717 | |
| 
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
 wenzelm parents: 
42379diff
changeset | 718 | val uncheck_typs = Syntax.typ_uncheck; | 
| 
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
 wenzelm parents: 
42379diff
changeset | 719 | val uncheck_terms = Syntax.term_uncheck; | 
| 
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
 wenzelm parents: 
42379diff
changeset | 720 | |
| 
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
 wenzelm parents: 
42379diff
changeset | 721 | |
| 
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
 wenzelm parents: 
42379diff
changeset | 722 | |
| 42242 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 723 | (** install operations **) | 
| 
39261908e12f
moved decode/parse operations to standard_syntax.ML;
 wenzelm parents: 
42241diff
changeset | 724 | |
| 42241 
dd8029f71e1c
separate module for standard implementation of inner syntax operations;
 wenzelm parents: diff
changeset | 725 | val _ = Syntax.install_operations | 
| 
dd8029f71e1c
separate module for standard implementation of inner syntax operations;
 wenzelm parents: diff
changeset | 726 |   {parse_sort = parse_sort,
 | 
| 
dd8029f71e1c
separate module for standard implementation of inner syntax operations;
 wenzelm parents: diff
changeset | 727 | parse_typ = parse_typ, | 
| 42281 | 728 | parse_term = parse_term false, | 
| 729 | parse_prop = parse_term true, | |
| 42241 
dd8029f71e1c
separate module for standard implementation of inner syntax operations;
 wenzelm parents: diff
changeset | 730 | unparse_sort = unparse_sort, | 
| 
dd8029f71e1c
separate module for standard implementation of inner syntax operations;
 wenzelm parents: diff
changeset | 731 | unparse_typ = unparse_typ, | 
| 42382 
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
 wenzelm parents: 
42379diff
changeset | 732 | unparse_term = unparse_term, | 
| 
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
 wenzelm parents: 
42379diff
changeset | 733 | check_typs = check_typs, | 
| 
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
 wenzelm parents: 
42379diff
changeset | 734 | check_terms = check_terms, | 
| 
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
 wenzelm parents: 
42379diff
changeset | 735 | check_props = check_props, | 
| 
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
 wenzelm parents: 
42379diff
changeset | 736 | uncheck_typs = uncheck_typs, | 
| 
dcd983ee2c29
provide structure Syntax early (before structure Type), back-patch check/uncheck later;
 wenzelm parents: 
42379diff
changeset | 737 | uncheck_terms = uncheck_terms}; | 
| 42241 
dd8029f71e1c
separate module for standard implementation of inner syntax operations;
 wenzelm parents: diff
changeset | 738 | |
| 
dd8029f71e1c
separate module for standard implementation of inner syntax operations;
 wenzelm parents: diff
changeset | 739 | end; |