src/Pure/Syntax/syntax_phases.ML
author wenzelm
Thu, 10 Nov 2011 22:54:15 +0100
changeset 45447 a97251eea458
parent 45445 41e641a870de
child 45448 018f8959c7a6
permissions -rw-r--r--
suppress irrelevant positions;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
42243
2f998ff67d0f renamed Standard_Syntax to Syntax_Phases;
wenzelm
parents: 42242
diff changeset
     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
2f998ff67d0f renamed Standard_Syntax to Syntax_Phases;
wenzelm
parents: 42242
diff changeset
     4
Main phases of inner syntax processing, with standard implementations
2f998ff67d0f renamed Standard_Syntax to Syntax_Phases;
wenzelm
parents: 42242
diff changeset
     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
2f998ff67d0f renamed Standard_Syntax to Syntax_Phases;
wenzelm
parents: 42242
diff changeset
     8
signature SYNTAX_PHASES =
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
     9
sig
45427
fca432074fb2 sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents: 45423
diff changeset
    10
  val decode_sort: term -> sort
fca432074fb2 sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents: 45423
diff changeset
    11
  val decode_typ: term -> typ
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
    12
  val decode_term: Proof.context ->
44736
c2a3f1c84179 bulk reports for improved message throughput;
wenzelm
parents: 44735
diff changeset
    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: 42241
diff changeset
    14
  val parse_ast_pattern: Proof.context -> string * string -> Ast.ast
42247
12fe41a92cd5 typed_print_translation: discontinued show_sorts argument;
wenzelm
parents: 42245
diff changeset
    15
  val term_of_typ: Proof.context -> typ -> term
45429
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
    16
  val print_checks: Proof.context -> unit
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
    17
  val typ_check: int -> string -> (Proof.context -> typ list -> typ list) ->
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
    18
    Context.generic -> Context.generic
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
    19
  val term_check: int -> string -> (Proof.context -> term list -> term list) ->
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
    20
    Context.generic -> Context.generic
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
    21
  val typ_uncheck: int -> string -> (Proof.context -> typ list -> typ list) ->
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
    22
    Context.generic -> Context.generic
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
    23
  val term_uncheck: int -> string -> (Proof.context -> term list -> term list) ->
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
    24
    Context.generic -> Context.generic
45444
ac069060e08a tuned signature;
wenzelm
parents: 45429
diff changeset
    25
  val typ_check': int -> string ->
ac069060e08a tuned signature;
wenzelm
parents: 45429
diff changeset
    26
    (typ list -> Proof.context -> (typ list * Proof.context) option) ->
ac069060e08a tuned signature;
wenzelm
parents: 45429
diff changeset
    27
    Context.generic -> Context.generic
ac069060e08a tuned signature;
wenzelm
parents: 45429
diff changeset
    28
  val term_check': int -> string ->
ac069060e08a tuned signature;
wenzelm
parents: 45429
diff changeset
    29
    (term list -> Proof.context -> (term list * Proof.context) option) ->
ac069060e08a tuned signature;
wenzelm
parents: 45429
diff changeset
    30
    Context.generic -> Context.generic
ac069060e08a tuned signature;
wenzelm
parents: 45429
diff changeset
    31
  val typ_uncheck': int -> string ->
ac069060e08a tuned signature;
wenzelm
parents: 45429
diff changeset
    32
    (typ list -> Proof.context -> (typ list * Proof.context) option) ->
ac069060e08a tuned signature;
wenzelm
parents: 45429
diff changeset
    33
    Context.generic -> Context.generic
ac069060e08a tuned signature;
wenzelm
parents: 45429
diff changeset
    34
  val term_uncheck': int -> string ->
ac069060e08a tuned signature;
wenzelm
parents: 45429
diff changeset
    35
    (term list -> Proof.context -> (term list * Proof.context) option) ->
ac069060e08a tuned signature;
wenzelm
parents: 45429
diff changeset
    36
    Context.generic -> Context.generic
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
    37
end
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
    38
42243
2f998ff67d0f renamed Standard_Syntax to Syntax_Phases;
wenzelm
parents: 42242
diff changeset
    39
structure Syntax_Phases: SYNTAX_PHASES =
42241
dd8029f71e1c separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff changeset
    40
struct
dd8029f71e1c separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff changeset
    41
42298
d622145603ee more accurate markup for syntax consts, notably binders which point back to the original logical entity;
wenzelm
parents: 42296
diff changeset
    42
(** markup logical entities **)
d622145603ee more accurate markup for syntax consts, notably binders which point back to the original logical entity;
wenzelm
parents: 42296
diff changeset
    43
d622145603ee more accurate markup for syntax consts, notably binders which point back to the original logical entity;
wenzelm
parents: 42296
diff changeset
    44
fun markup_class ctxt c =
42379
26f64dddf2c6 tuned signature;
wenzelm
parents: 42360
diff changeset
    45
  [Name_Space.markup (Type.class_space (Proof_Context.tsig_of ctxt)) c];
42298
d622145603ee more accurate markup for syntax consts, notably binders which point back to the original logical entity;
wenzelm
parents: 42296
diff changeset
    46
d622145603ee more accurate markup for syntax consts, notably binders which point back to the original logical entity;
wenzelm
parents: 42296
diff changeset
    47
fun markup_type ctxt c =
42379
26f64dddf2c6 tuned signature;
wenzelm
parents: 42360
diff changeset
    48
  [Name_Space.markup (Type.type_space (Proof_Context.tsig_of ctxt)) c];
42298
d622145603ee more accurate markup for syntax consts, notably binders which point back to the original logical entity;
wenzelm
parents: 42296
diff changeset
    49
d622145603ee more accurate markup for syntax consts, notably binders which point back to the original logical entity;
wenzelm
parents: 42296
diff changeset
    50
fun markup_const ctxt c =
42379
26f64dddf2c6 tuned signature;
wenzelm
parents: 42360
diff changeset
    51
  [Name_Space.markup (Consts.space_of (Proof_Context.consts_of ctxt)) c];
42298
d622145603ee more accurate markup for syntax consts, notably binders which point back to the original logical entity;
wenzelm
parents: 42296
diff changeset
    52
d622145603ee more accurate markup for syntax consts, notably binders which point back to the original logical entity;
wenzelm
parents: 42296
diff changeset
    53
fun markup_free ctxt x =
d622145603ee more accurate markup for syntax consts, notably binders which point back to the original logical entity;
wenzelm
parents: 42296
diff changeset
    54
  [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: 42488
diff changeset
    55
  (if Variable.is_body ctxt orelse Variable.is_fixed ctxt x
01430341fc79 more informative markup for fixed variables (via name space entry);
wenzelm
parents: 42488
diff changeset
    56
   then [Variable.markup_fixed ctxt x]
01430341fc79 more informative markup for fixed variables (via name space entry);
wenzelm
parents: 42488
diff changeset
    57
   else []);
42298
d622145603ee more accurate markup for syntax consts, notably binders which point back to the original logical entity;
wenzelm
parents: 42296
diff changeset
    58
d622145603ee more accurate markup for syntax consts, notably binders which point back to the original logical entity;
wenzelm
parents: 42296
diff changeset
    59
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: 42296
diff changeset
    60
45412
7797f5351ec4 entity markup for bound variables;
wenzelm
parents: 44736
diff changeset
    61
fun markup_bound def ps (name, id) =
7797f5351ec4 entity markup for bound variables;
wenzelm
parents: 44736
diff changeset
    62
  let val entity = Markup.entity "bound variable" name in
7797f5351ec4 entity markup for bound variables;
wenzelm
parents: 44736
diff changeset
    63
    Markup.bound ::
7797f5351ec4 entity markup for bound variables;
wenzelm
parents: 44736
diff changeset
    64
      map (fn pos => Markup.properties (Position.entity_properties_of def id pos) entity) ps
7797f5351ec4 entity markup for bound variables;
wenzelm
parents: 44736
diff changeset
    65
  end;
42298
d622145603ee more accurate markup for syntax consts, notably binders which point back to the original logical entity;
wenzelm
parents: 42296
diff changeset
    66
d622145603ee more accurate markup for syntax consts, notably binders which point back to the original logical entity;
wenzelm
parents: 42296
diff changeset
    67
fun markup_entity ctxt c =
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 42359
diff changeset
    68
  (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: 42296
diff changeset
    69
    SOME "" => []
d622145603ee more accurate markup for syntax consts, notably binders which point back to the original logical entity;
wenzelm
parents: 42296
diff changeset
    70
  | 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: 42296
diff changeset
    71
  | NONE => c |> Lexicon.unmark
d622145603ee more accurate markup for syntax consts, notably binders which point back to the original logical entity;
wenzelm
parents: 42296
diff changeset
    72
     {case_class = markup_class ctxt,
d622145603ee more accurate markup for syntax consts, notably binders which point back to the original logical entity;
wenzelm
parents: 42296
diff changeset
    73
      case_type = markup_type ctxt,
d622145603ee more accurate markup for syntax consts, notably binders which point back to the original logical entity;
wenzelm
parents: 42296
diff changeset
    74
      case_const = markup_const ctxt,
d622145603ee more accurate markup for syntax consts, notably binders which point back to the original logical entity;
wenzelm
parents: 42296
diff changeset
    75
      case_fixed = markup_free ctxt,
d622145603ee more accurate markup for syntax consts, notably binders which point back to the original logical entity;
wenzelm
parents: 42296
diff changeset
    76
      case_default = K []});
d622145603ee more accurate markup for syntax consts, notably binders which point back to the original logical entity;
wenzelm
parents: 42296
diff changeset
    77
d622145603ee more accurate markup for syntax consts, notably binders which point back to the original logical entity;
wenzelm
parents: 42296
diff changeset
    78
d622145603ee more accurate markup for syntax consts, notably binders which point back to the original logical entity;
wenzelm
parents: 42296
diff changeset
    79
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
    80
(** decode parse trees **)
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
    81
45427
fca432074fb2 sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents: 45423
diff changeset
    82
(* decode_sort *)
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
    83
45427
fca432074fb2 sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents: 45423
diff changeset
    84
fun decode_sort tm =
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
    85
  let
45427
fca432074fb2 sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents: 45423
diff changeset
    86
    fun err () = raise TERM ("decode_sort: bad encoding of classes", [tm]);
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
    87
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
    88
    fun class s = Lexicon.unmark_class s handle Fail _ => err ();
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
    89
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
    90
    fun classes (Const (s, _)) = [class s]
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
    91
      | classes (Const ("_classes", _) $ Const (s, _) $ cs) = class s :: classes cs
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
    92
      | classes _ = err ();
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
    93
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
    94
    fun sort (Const ("_topsort", _)) = []
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
    95
      | sort (Const (s, _)) = [class s]
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
    96
      | sort (Const ("_sort", _) $ cs) = classes cs
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
    97
      | sort _ = err ();
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
    98
  in sort tm end;
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
    99
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   100
45427
fca432074fb2 sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents: 45423
diff changeset
   101
(* decode_typ *)
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   102
45427
fca432074fb2 sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents: 45423
diff changeset
   103
fun decode_typ tm =
fca432074fb2 sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents: 45423
diff changeset
   104
  let
fca432074fb2 sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents: 45423
diff changeset
   105
    fun err () = raise TERM ("decode_typ: bad encoding of type", [tm]);
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   106
45427
fca432074fb2 sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents: 45423
diff changeset
   107
    fun typ (Free (x, _)) = TFree (x, dummyS)
fca432074fb2 sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents: 45423
diff changeset
   108
      | typ (Var (xi, _)) = TVar (xi, dummyS)
fca432074fb2 sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents: 45423
diff changeset
   109
      | typ (Const ("_tfree",_) $ (t as Free _)) = typ t
fca432074fb2 sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents: 45423
diff changeset
   110
      | typ (Const ("_tvar",_) $ (t as Var _)) = typ t
fca432074fb2 sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents: 45423
diff changeset
   111
      | typ (Const ("_ofsort", _) $ Free (x, _) $ s) = TFree (x, decode_sort s)
fca432074fb2 sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents: 45423
diff changeset
   112
      | typ (Const ("_ofsort", _) $ (Const ("_tfree",_) $ Free (x, _)) $ s) =
fca432074fb2 sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents: 45423
diff changeset
   113
          TFree (x, decode_sort s)
fca432074fb2 sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents: 45423
diff changeset
   114
      | typ (Const ("_ofsort", _) $ Var (xi, _) $ s) = TVar (xi, decode_sort s)
fca432074fb2 sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents: 45423
diff changeset
   115
      | typ (Const ("_ofsort", _) $ (Const ("_tvar",_) $ Var (xi, _)) $ s) =
fca432074fb2 sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents: 45423
diff changeset
   116
          TVar (xi, decode_sort s)
fca432074fb2 sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents: 45423
diff changeset
   117
      | typ (Const ("_dummy_ofsort", _) $ s) = TFree ("'_dummy_", decode_sort s)
fca432074fb2 sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents: 45423
diff changeset
   118
      | typ t =
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   119
          let
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   120
            val (head, args) = Term.strip_comb t;
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   121
            val a =
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   122
              (case head of
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   123
                Const (c, _) => (Lexicon.unmark_type c handle Fail _ => err ())
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   124
              | _ => err ());
45427
fca432074fb2 sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents: 45423
diff changeset
   125
          in Type (a, map typ args) end;
fca432074fb2 sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents: 45423
diff changeset
   126
  in typ tm end;
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   127
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   128
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   129
(* parsetree_to_ast *)
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   130
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   131
fun parsetree_to_ast ctxt constrain_pos trf parsetree =
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   132
  let
44736
c2a3f1c84179 bulk reports for improved message throughput;
wenzelm
parents: 44735
diff changeset
   133
    val reports = Unsynchronized.ref ([]: Position.report list);
44735
66862d02678c tuned signature;
wenzelm
parents: 44564
diff changeset
   134
    fun report pos = Position.store_reports reports [pos];
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   135
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   136
    fun trans a args =
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   137
      (case trf a of
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   138
        NONE => Ast.mk_appl (Ast.Constant a) args
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   139
      | SOME f => f ctxt args);
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   140
42282
4fa41e068ff0 report literal tokens according to parsetree head;
wenzelm
parents: 42281
diff changeset
   141
    fun asts_of (Parser.Node ("_class_name", [Parser.Tip tok])) =
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   142
          let
42470
cc78b0ed0fad more precise error positions;
wenzelm
parents: 42410
diff changeset
   143
            val pos = Lexicon.pos_of_token tok;
cc78b0ed0fad more precise error positions;
wenzelm
parents: 42410
diff changeset
   144
            val c = Proof_Context.read_class ctxt (Lexicon.str_of_token tok)
cc78b0ed0fad more precise error positions;
wenzelm
parents: 42410
diff changeset
   145
              handle ERROR msg => error (msg ^ Position.str_of pos);
cc78b0ed0fad more precise error positions;
wenzelm
parents: 42410
diff changeset
   146
            val _ = report pos (markup_class ctxt) c;
42282
4fa41e068ff0 report literal tokens according to parsetree head;
wenzelm
parents: 42281
diff changeset
   147
          in [Ast.Constant (Lexicon.mark_class c)] end
4fa41e068ff0 report literal tokens according to parsetree head;
wenzelm
parents: 42281
diff changeset
   148
      | asts_of (Parser.Node ("_type_name", [Parser.Tip tok])) =
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   149
          let
42470
cc78b0ed0fad more precise error positions;
wenzelm
parents: 42410
diff changeset
   150
            val pos = Lexicon.pos_of_token tok;
cc78b0ed0fad more precise error positions;
wenzelm
parents: 42410
diff changeset
   151
            val Type (c, _) =
cc78b0ed0fad more precise error positions;
wenzelm
parents: 42410
diff changeset
   152
              Proof_Context.read_type_name_proper ctxt false (Lexicon.str_of_token tok)
cc78b0ed0fad more precise error positions;
wenzelm
parents: 42410
diff changeset
   153
                handle ERROR msg => error (msg ^ Position.str_of pos);
cc78b0ed0fad more precise error positions;
wenzelm
parents: 42410
diff changeset
   154
            val _ = report pos (markup_type ctxt) c;
42282
4fa41e068ff0 report literal tokens according to parsetree head;
wenzelm
parents: 42281
diff changeset
   155
          in [Ast.Constant (Lexicon.mark_type c)] end
42410
16bc5564535f less bulky "_position", for improved readability of parse trees;
wenzelm
parents: 42408
diff changeset
   156
      | asts_of (Parser.Node ("_position", [pt as Parser.Tip tok])) =
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   157
          if constrain_pos then
45447
a97251eea458 suppress irrelevant positions;
wenzelm
parents: 45445
diff changeset
   158
            let val pos = Lexicon.pos_of_token tok in
a97251eea458 suppress irrelevant positions;
wenzelm
parents: 45445
diff changeset
   159
              if Position.is_reported pos then
a97251eea458 suppress irrelevant positions;
wenzelm
parents: 45445
diff changeset
   160
                [Ast.Appl [Ast.Constant "_constrain", ast_of pt,
a97251eea458 suppress irrelevant positions;
wenzelm
parents: 45445
diff changeset
   161
                  Ast.Variable (Term_Position.encode pos)]]
a97251eea458 suppress irrelevant positions;
wenzelm
parents: 45445
diff changeset
   162
              else [ast_of pt]
a97251eea458 suppress irrelevant positions;
wenzelm
parents: 45445
diff changeset
   163
            end
42282
4fa41e068ff0 report literal tokens according to parsetree head;
wenzelm
parents: 42281
diff changeset
   164
          else [ast_of pt]
4fa41e068ff0 report literal tokens according to parsetree head;
wenzelm
parents: 42281
diff changeset
   165
      | asts_of (Parser.Node (a, pts)) =
4fa41e068ff0 report literal tokens according to parsetree head;
wenzelm
parents: 42281
diff changeset
   166
          let
4fa41e068ff0 report literal tokens according to parsetree head;
wenzelm
parents: 42281
diff changeset
   167
            val _ = pts |> List.app
4fa41e068ff0 report literal tokens according to parsetree head;
wenzelm
parents: 42281
diff changeset
   168
              (fn Parser.Node _ => () | Parser.Tip tok =>
4fa41e068ff0 report literal tokens according to parsetree head;
wenzelm
parents: 42281
diff changeset
   169
                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: 42296
diff changeset
   170
                else report (Lexicon.pos_of_token tok) (markup_entity ctxt) a);
42282
4fa41e068ff0 report literal tokens according to parsetree head;
wenzelm
parents: 42281
diff changeset
   171
          in [trans a (maps asts_of pts)] end
4fa41e068ff0 report literal tokens according to parsetree head;
wenzelm
parents: 42281
diff changeset
   172
      | asts_of (Parser.Tip tok) =
4fa41e068ff0 report literal tokens according to parsetree head;
wenzelm
parents: 42281
diff changeset
   173
          if Lexicon.valued_token tok
4fa41e068ff0 report literal tokens according to parsetree head;
wenzelm
parents: 42281
diff changeset
   174
          then [Ast.Variable (Lexicon.str_of_token tok)]
4fa41e068ff0 report literal tokens according to parsetree head;
wenzelm
parents: 42281
diff changeset
   175
          else []
4fa41e068ff0 report literal tokens according to parsetree head;
wenzelm
parents: 42281
diff changeset
   176
4fa41e068ff0 report literal tokens according to parsetree head;
wenzelm
parents: 42281
diff changeset
   177
    and ast_of pt =
4fa41e068ff0 report literal tokens according to parsetree head;
wenzelm
parents: 42281
diff changeset
   178
      (case asts_of pt of
4fa41e068ff0 report literal tokens according to parsetree head;
wenzelm
parents: 42281
diff changeset
   179
        [ast] => ast
4fa41e068ff0 report literal tokens according to parsetree head;
wenzelm
parents: 42281
diff changeset
   180
      | asts => raise Ast.AST ("parsetree_to_ast: malformed parsetree", asts));
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   181
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   182
    val ast = Exn.interruptible_capture ast_of parsetree;
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   183
  in (! reports, ast) end;
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   184
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   185
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   186
(* ast_to_term *)
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   187
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   188
fun ast_to_term ctxt trf =
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   189
  let
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   190
    fun trans a args =
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   191
      (case trf a of
42476
d0bc1268ef09 clarified auxiliary structure Lexicon.Syntax;
wenzelm
parents: 42470
diff changeset
   192
        NONE => Term.list_comb (Syntax.const a, args)
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   193
      | SOME f => f ctxt args);
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   194
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   195
    fun term_of (Ast.Constant a) = trans a []
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   196
      | term_of (Ast.Variable x) = Lexicon.read_var x
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   197
      | term_of (Ast.Appl (Ast.Constant a :: (asts as _ :: _))) =
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   198
          trans a (map term_of asts)
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   199
      | term_of (Ast.Appl (ast :: (asts as _ :: _))) =
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   200
          Term.list_comb (term_of ast, map term_of asts)
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   201
      | term_of (ast as Ast.Appl _) = raise Ast.AST ("ast_to_term: malformed ast", [ast]);
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   202
  in term_of end;
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   203
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   204
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   205
(* decode_term -- transform parse tree into raw term *)
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   206
44736
c2a3f1c84179 bulk reports for improved message throughput;
wenzelm
parents: 44735
diff changeset
   207
fun decode_term _ (result as (_: Position.report list, Exn.Exn _)) = result
43761
e72ba84ae58f tuned signature -- corresponding to Scala version;
wenzelm
parents: 43731
diff changeset
   208
  | decode_term ctxt (reports0, Exn.Res tm) =
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   209
      let
42250
cc5ac538f991 eliminated odd object-oriented type_context/term_context;
wenzelm
parents: 42249
diff changeset
   210
        fun get_const a =
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 42359
diff changeset
   211
          ((true, #1 (Term.dest_Const (Proof_Context.read_const_proper ctxt false a)))
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 42359
diff changeset
   212
            handle ERROR _ => (false, Consts.intern (Proof_Context.consts_of ctxt) a));
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 42359
diff changeset
   213
        val get_free = Proof_Context.intern_skolem ctxt;
42250
cc5ac538f991 eliminated odd object-oriented type_context/term_context;
wenzelm
parents: 42249
diff changeset
   214
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   215
        val reports = Unsynchronized.ref reports0;
44735
66862d02678c tuned signature;
wenzelm
parents: 44564
diff changeset
   216
        fun report ps = Position.store_reports reports ps;
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   217
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   218
        fun decode ps qs bs (Const ("_constrain", _) $ t $ typ) =
42264
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents: 42255
diff changeset
   219
              (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: 45444
diff changeset
   220
                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: 45423
diff changeset
   221
              | NONE => Type.constraint (decode_typ typ) (decode ps qs bs t))
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   222
          | decode ps qs bs (Const ("_constrainAbs", _) $ t $ typ) =
42264
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents: 42255
diff changeset
   223
              (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: 45444
diff changeset
   224
                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: 45423
diff changeset
   225
              | NONE => Type.constraint (decode_typ typ --> dummyT) (decode ps qs bs t))
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   226
          | decode _ qs bs (Abs (x, T, t)) =
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   227
              let
45412
7797f5351ec4 entity markup for bound variables;
wenzelm
parents: 44736
diff changeset
   228
                val id = serial ();
7797f5351ec4 entity markup for bound variables;
wenzelm
parents: 44736
diff changeset
   229
                val _ = report qs (markup_bound true qs) (x, id);
7797f5351ec4 entity markup for bound variables;
wenzelm
parents: 44736
diff changeset
   230
              in Abs (x, T, decode [] [] ((qs, (x, id)) :: bs) t) end
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   231
          | decode _ _ bs (t $ u) = decode [] [] bs t $ decode [] [] bs u
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   232
          | decode ps _ _ (Const (a, T)) =
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   233
              (case try Lexicon.unmark_fixed a of
42282
4fa41e068ff0 report literal tokens according to parsetree head;
wenzelm
parents: 42281
diff changeset
   234
                SOME x => (report ps (markup_free ctxt) x; Free (x, T))
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   235
              | NONE =>
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   236
                  let
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   237
                    val c =
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   238
                      (case try Lexicon.unmark_const a of
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   239
                        SOME c => c
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   240
                      | NONE => snd (get_const a));
42282
4fa41e068ff0 report literal tokens according to parsetree head;
wenzelm
parents: 42281
diff changeset
   241
                    val _ = report ps (markup_const ctxt) c;
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   242
                  in Const (c, T) end)
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   243
          | decode ps _ _ (Free (a, T)) =
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   244
              (case (get_free a, get_const a) of
42282
4fa41e068ff0 report literal tokens according to parsetree head;
wenzelm
parents: 42281
diff changeset
   245
                (SOME x, _) => (report ps (markup_free ctxt) x; Free (x, T))
4fa41e068ff0 report literal tokens according to parsetree head;
wenzelm
parents: 42281
diff changeset
   246
              | (_, (true, c)) => (report ps (markup_const ctxt) c; Const (c, T))
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   247
              | (_, (false, c)) =>
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   248
                  if Long_Name.is_qualified c
42282
4fa41e068ff0 report literal tokens according to parsetree head;
wenzelm
parents: 42281
diff changeset
   249
                  then (report ps (markup_const ctxt) c; Const (c, T))
4fa41e068ff0 report literal tokens according to parsetree head;
wenzelm
parents: 42281
diff changeset
   250
                  else (report ps (markup_free ctxt) c; Free (c, T)))
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   251
          | decode ps _ _ (Var (xi, T)) = (report ps markup_var xi; Var (xi, T))
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   252
          | decode ps _ bs (t as Bound i) =
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   253
              (case try (nth bs) i of
45412
7797f5351ec4 entity markup for bound variables;
wenzelm
parents: 44736
diff changeset
   254
                SOME (qs, (x, id)) => (report ps (markup_bound false qs) (x, id); t)
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   255
              | NONE => t);
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   256
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   257
        val tm' = Exn.interruptible_capture (fn () => decode [] [] [] tm) ();
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   258
      in (! reports, tm') end;
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   259
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   260
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   261
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   262
(** parse **)
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   263
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   264
(* results *)
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   265
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   266
fun ambiguity_msg pos = "Parse error: ambiguous syntax" ^ Position.str_of pos;
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   267
43761
e72ba84ae58f tuned signature -- corresponding to Scala version;
wenzelm
parents: 43731
diff changeset
   268
fun proper_results results = map_filter (fn (y, Exn.Res x) => SOME (y, x) | _ => NONE) results;
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   269
fun failed_results results = map_filter (fn (y, Exn.Exn e) => SOME (y, e) | _ => NONE) results;
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   270
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   271
fun report_result ctxt pos results =
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   272
  (case (proper_results results, failed_results results) of
44736
c2a3f1c84179 bulk reports for improved message throughput;
wenzelm
parents: 44735
diff changeset
   273
    ([], (reports, exn) :: _) => (Context_Position.reports ctxt reports; reraise exn)
c2a3f1c84179 bulk reports for improved message throughput;
wenzelm
parents: 44735
diff changeset
   274
  | ([(reports, x)], _) => (Context_Position.reports ctxt reports; x)
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   275
  | _ => error (ambiguity_msg pos));
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   276
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   277
42249
12a073670584 simplified standard parse/unparse;
wenzelm
parents: 42248
diff changeset
   278
(* parse raw asts *)
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   279
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   280
fun parse_asts ctxt raw root (syms, pos) =
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   281
  let
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 42359
diff changeset
   282
    val syn = Proof_Context.syn_of ctxt;
42253
c539d3c9c023 more abstract syntax translation;
wenzelm
parents: 42252
diff changeset
   283
    val ast_tr = Syntax.parse_ast_translation syn;
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   284
42251
050cc12dd985 explicit Syntax.tokenize, Syntax.parse;
wenzelm
parents: 42250
diff changeset
   285
    val toks = Syntax.tokenize syn raw syms;
44736
c2a3f1c84179 bulk reports for improved message throughput;
wenzelm
parents: 44735
diff changeset
   286
    val _ = Context_Position.reports ctxt (map Lexicon.report_of_token toks);
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   287
42251
050cc12dd985 explicit Syntax.tokenize, Syntax.parse;
wenzelm
parents: 42250
diff changeset
   288
    val pts = Syntax.parse ctxt syn root (filter Lexicon.is_proper toks)
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   289
      handle ERROR msg =>
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   290
        error (msg ^
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   291
          implode (map (Markup.markup Markup.report o Lexicon.reported_token_range ctxt) toks));
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   292
    val len = length pts;
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   293
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   294
    val limit = Config.get ctxt Syntax.ambiguity_limit;
44069
d7c7ec248ef0 make syntax ambiguity warnings a config option
kleing
parents: 43794
diff changeset
   295
    val warnings = Config.get ctxt Syntax.ambiguity_warnings;
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   296
    val _ =
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   297
      if len <= Config.get ctxt Syntax.ambiguity_level then ()
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   298
      else if not (Config.get ctxt Syntax.ambiguity_enabled) then error (ambiguity_msg pos)
44069
d7c7ec248ef0 make syntax ambiguity warnings a config option
kleing
parents: 43794
diff changeset
   299
      else if warnings then
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   300
        (Context_Position.if_visible ctxt warning (cat_lines
44564
96ba83710946 tuned positions of ambiguity messages -- less intrusive in IDE view;
wenzelm
parents: 44069
diff changeset
   301
          (("Ambiguous input" ^ Position.str_of (Position.reset_range pos) ^
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   302
            "\nproduces " ^ string_of_int len ^ " parse trees" ^
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   303
            (if len <= limit then "" else " (" ^ string_of_int limit ^ " displayed)") ^ ":") ::
44069
d7c7ec248ef0 make syntax ambiguity warnings a config option
kleing
parents: 43794
diff changeset
   304
            map (Pretty.string_of o Parser.pretty_parsetree) (take limit pts))))
d7c7ec248ef0 make syntax ambiguity warnings a config option
kleing
parents: 43794
diff changeset
   305
      else ();
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   306
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   307
    val constrain_pos = not raw andalso Config.get ctxt Syntax.positions;
42253
c539d3c9c023 more abstract syntax translation;
wenzelm
parents: 42252
diff changeset
   308
    val parsetree_to_ast = parsetree_to_ast ctxt constrain_pos ast_tr;
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   309
  in map parsetree_to_ast pts end;
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   310
42249
12a073670584 simplified standard parse/unparse;
wenzelm
parents: 42248
diff changeset
   311
fun parse_raw ctxt root input =
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   312
  let
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 42359
diff changeset
   313
    val syn = Proof_Context.syn_of ctxt;
42253
c539d3c9c023 more abstract syntax translation;
wenzelm
parents: 42252
diff changeset
   314
    val tr = Syntax.parse_translation syn;
42255
097c93dcd877 eliminated slightly odd Syntax.rep_syntax;
wenzelm
parents: 42254
diff changeset
   315
    val parse_rules = Syntax.parse_rules syn;
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   316
  in
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   317
    parse_asts ctxt false root input
42255
097c93dcd877 eliminated slightly odd Syntax.rep_syntax;
wenzelm
parents: 42254
diff changeset
   318
    |> (map o apsnd o Exn.maps_result)
097c93dcd877 eliminated slightly odd Syntax.rep_syntax;
wenzelm
parents: 42254
diff changeset
   319
        (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: 42241
diff changeset
   320
  end;
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   321
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   322
42249
12a073670584 simplified standard parse/unparse;
wenzelm
parents: 42248
diff changeset
   323
(* parse logical entities *)
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   324
42241
dd8029f71e1c separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff changeset
   325
fun parse_failed ctxt pos msg kind =
dd8029f71e1c separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff changeset
   326
  cat_error msg ("Failed to parse " ^ kind ^
dd8029f71e1c separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff changeset
   327
    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
   328
43731
70072780e095 inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents: 43552
diff changeset
   329
fun parse_sort ctxt =
70072780e095 inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents: 43552
diff changeset
   330
  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: 43552
diff changeset
   331
    (fn (syms, pos) =>
42249
12a073670584 simplified standard parse/unparse;
wenzelm
parents: 42248
diff changeset
   332
      parse_raw ctxt "sort" (syms, pos)
12a073670584 simplified standard parse/unparse;
wenzelm
parents: 42248
diff changeset
   333
      |> report_result ctxt pos
45427
fca432074fb2 sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents: 45423
diff changeset
   334
      |> decode_sort
43731
70072780e095 inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents: 43552
diff changeset
   335
      |> Type.minimize_sort (Proof_Context.tsig_of ctxt)
70072780e095 inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents: 43552
diff changeset
   336
      handle ERROR msg => parse_failed ctxt pos msg "sort");
42241
dd8029f71e1c separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff changeset
   337
43731
70072780e095 inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents: 43552
diff changeset
   338
fun parse_typ ctxt =
70072780e095 inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents: 43552
diff changeset
   339
  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: 43552
diff changeset
   340
    (fn (syms, pos) =>
42249
12a073670584 simplified standard parse/unparse;
wenzelm
parents: 42248
diff changeset
   341
      parse_raw ctxt "type" (syms, pos)
12a073670584 simplified standard parse/unparse;
wenzelm
parents: 42248
diff changeset
   342
      |> report_result ctxt pos
45427
fca432074fb2 sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents: 45423
diff changeset
   343
      |> decode_typ
43731
70072780e095 inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents: 43552
diff changeset
   344
      handle ERROR msg => parse_failed ctxt pos msg "type");
42241
dd8029f71e1c separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff changeset
   345
43731
70072780e095 inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents: 43552
diff changeset
   346
fun parse_term is_prop ctxt =
42241
dd8029f71e1c separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff changeset
   347
  let
42281
69d4543811d0 simplified read_term vs. read_prop;
wenzelm
parents: 42267
diff changeset
   348
    val (markup, kind, root, constrain) =
69d4543811d0 simplified read_term vs. read_prop;
wenzelm
parents: 42267
diff changeset
   349
      if is_prop
69d4543811d0 simplified read_term vs. read_prop;
wenzelm
parents: 42267
diff changeset
   350
      then (Markup.prop, "proposition", "prop", Type.constraint propT)
42288
2074b31650e6 discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
wenzelm
parents: 42284
diff changeset
   351
      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: 43552
diff changeset
   352
    val decode = constrain o Term_XML.Decode.term;
42249
12a073670584 simplified standard parse/unparse;
wenzelm
parents: 42248
diff changeset
   353
  in
43731
70072780e095 inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents: 43552
diff changeset
   354
    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: 43552
diff changeset
   355
      (fn (syms, pos) =>
70072780e095 inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents: 43552
diff changeset
   356
        let
70072780e095 inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents: 43552
diff changeset
   357
          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: 43552
diff changeset
   358
          val ambiguity = length (proper_results results);
42249
12a073670584 simplified standard parse/unparse;
wenzelm
parents: 42248
diff changeset
   359
43731
70072780e095 inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents: 43552
diff changeset
   360
          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: 43552
diff changeset
   361
          val limit = Config.get ctxt Syntax.ambiguity_limit;
44069
d7c7ec248ef0 make syntax ambiguity warnings a config option
kleing
parents: 43794
diff changeset
   362
          val warnings = Config.get ctxt Syntax.ambiguity_warnings;
42249
12a073670584 simplified standard parse/unparse;
wenzelm
parents: 42248
diff changeset
   363
43731
70072780e095 inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents: 43552
diff changeset
   364
          val ambig_msg =
70072780e095 inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents: 43552
diff changeset
   365
            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: 43552
diff changeset
   366
              ["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: 43552
diff changeset
   367
              \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: 43552
diff changeset
   368
            else [];
70072780e095 inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents: 43552
diff changeset
   369
70072780e095 inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents: 43552
diff changeset
   370
          (*brute-force disambiguation via type-inference*)
43761
e72ba84ae58f tuned signature -- corresponding to Scala version;
wenzelm
parents: 43731
diff changeset
   371
          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: 43552
diff changeset
   372
            handle exn as ERROR _ => Exn.Exn exn;
42249
12a073670584 simplified standard parse/unparse;
wenzelm
parents: 42248
diff changeset
   373
43731
70072780e095 inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents: 43552
diff changeset
   374
          val results' =
70072780e095 inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents: 43552
diff changeset
   375
            if ambiguity > 1 then
70072780e095 inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents: 43552
diff changeset
   376
              (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: 43552
diff changeset
   377
                check results
70072780e095 inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents: 43552
diff changeset
   378
            else results;
70072780e095 inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents: 43552
diff changeset
   379
          val reports' = fst (hd results');
42249
12a073670584 simplified standard parse/unparse;
wenzelm
parents: 42248
diff changeset
   380
43731
70072780e095 inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents: 43552
diff changeset
   381
          val errs = map snd (failed_results results');
70072780e095 inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents: 43552
diff changeset
   382
          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: 43552
diff changeset
   383
          val len = length checked;
42249
12a073670584 simplified standard parse/unparse;
wenzelm
parents: 42248
diff changeset
   384
43731
70072780e095 inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents: 43552
diff changeset
   385
          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: 43552
diff changeset
   386
        in
70072780e095 inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents: 43552
diff changeset
   387
          if len = 0 then
70072780e095 inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents: 43552
diff changeset
   388
            report_result ctxt pos
70072780e095 inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents: 43552
diff changeset
   389
              [(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: 43552
diff changeset
   390
          else if len = 1 then
44069
d7c7ec248ef0 make syntax ambiguity warnings a config option
kleing
parents: 43794
diff changeset
   391
            (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: 43552
diff changeset
   392
              Context_Position.if_visible ctxt warning
44564
96ba83710946 tuned positions of ambiguity messages -- less intrusive in IDE view;
wenzelm
parents: 44069
diff changeset
   393
                ("Fortunately, only one parse tree is type correct" ^
96ba83710946 tuned positions of ambiguity messages -- less intrusive in IDE view;
wenzelm
parents: 44069
diff changeset
   394
                Position.str_of (Position.reset_range pos) ^
96ba83710946 tuned positions of ambiguity messages -- less intrusive in IDE view;
wenzelm
parents: 44069
diff changeset
   395
                ",\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: 43552
diff changeset
   396
            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: 43552
diff changeset
   397
          else
70072780e095 inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents: 43552
diff changeset
   398
            report_result ctxt pos
70072780e095 inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents: 43552
diff changeset
   399
              [(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: 43552
diff changeset
   400
                (("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: 43552
diff changeset
   401
                  (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: 43552
diff changeset
   402
                  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: 43552
diff changeset
   403
        end handle ERROR msg => parse_failed ctxt pos msg kind)
42249
12a073670584 simplified standard parse/unparse;
wenzelm
parents: 42248
diff changeset
   404
  end;
42241
dd8029f71e1c separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff changeset
   405
dd8029f71e1c separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff changeset
   406
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   407
(* parse_ast_pattern *)
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   408
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   409
fun parse_ast_pattern ctxt (root, str) =
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   410
  let
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 42359
diff changeset
   411
    val syn = Proof_Context.syn_of ctxt;
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   412
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   413
    fun constify (ast as Ast.Constant _) = ast
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   414
      | 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: 42296
diff changeset
   415
          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: 42296
diff changeset
   416
          then Ast.Constant x
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   417
          else ast
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   418
      | constify (Ast.Appl asts) = Ast.Appl (map constify asts);
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   419
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   420
    val (syms, pos) = Syntax.read_token str;
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   421
  in
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   422
    parse_asts ctxt true root (syms, pos)
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   423
    |> report_result ctxt pos
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   424
    |> constify
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   425
  end;
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   426
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   427
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   428
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   429
(** encode parse trees **)
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   430
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   431
(* term_of_sort *)
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   432
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   433
fun term_of_sort S =
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   434
  let
42476
d0bc1268ef09 clarified auxiliary structure Lexicon.Syntax;
wenzelm
parents: 42470
diff changeset
   435
    val class = Syntax.const o Lexicon.mark_class;
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   436
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   437
    fun classes [c] = class c
42476
d0bc1268ef09 clarified auxiliary structure Lexicon.Syntax;
wenzelm
parents: 42470
diff changeset
   438
      | classes (c :: cs) = Syntax.const "_classes" $ class c $ classes cs;
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   439
  in
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   440
    (case S of
42476
d0bc1268ef09 clarified auxiliary structure Lexicon.Syntax;
wenzelm
parents: 42470
diff changeset
   441
      [] => Syntax.const "_topsort"
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   442
    | [c] => class c
42476
d0bc1268ef09 clarified auxiliary structure Lexicon.Syntax;
wenzelm
parents: 42470
diff changeset
   443
    | cs => Syntax.const "_sort" $ classes cs)
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   444
  end;
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   445
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   446
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   447
(* term_of_typ *)
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   448
42247
12fe41a92cd5 typed_print_translation: discontinued show_sorts argument;
wenzelm
parents: 42245
diff changeset
   449
fun term_of_typ ctxt ty =
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   450
  let
42247
12fe41a92cd5 typed_print_translation: discontinued show_sorts argument;
wenzelm
parents: 42245
diff changeset
   451
    val show_sorts = Config.get ctxt show_sorts;
12fe41a92cd5 typed_print_translation: discontinued show_sorts argument;
wenzelm
parents: 42245
diff changeset
   452
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   453
    fun of_sort t S =
42476
d0bc1268ef09 clarified auxiliary structure Lexicon.Syntax;
wenzelm
parents: 42470
diff changeset
   454
      if show_sorts then Syntax.const "_ofsort" $ t $ term_of_sort S
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   455
      else t;
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   456
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   457
    fun term_of (Type (a, Ts)) =
42476
d0bc1268ef09 clarified auxiliary structure Lexicon.Syntax;
wenzelm
parents: 42470
diff changeset
   458
          Term.list_comb (Syntax.const (Lexicon.mark_type a), map term_of Ts)
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   459
      | term_of (TFree (x, S)) =
42290
b1f544c84040 discontinued special treatment of structure Lexicon;
wenzelm
parents: 42289
diff changeset
   460
          if is_some (Term_Position.decode x) then Syntax.free x
42476
d0bc1268ef09 clarified auxiliary structure Lexicon.Syntax;
wenzelm
parents: 42470
diff changeset
   461
          else of_sort (Syntax.const "_tfree" $ Syntax.free x) S
d0bc1268ef09 clarified auxiliary structure Lexicon.Syntax;
wenzelm
parents: 42470
diff changeset
   462
      | term_of (TVar (xi, S)) = of_sort (Syntax.const "_tvar" $ Syntax.var xi) S;
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   463
  in term_of ty end;
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   464
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   465
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   466
(* simple_ast_of *)
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   467
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   468
fun simple_ast_of ctxt =
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   469
  let
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   470
    val tune_var = if Config.get ctxt show_question_marks then I else unprefix "?";
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   471
    fun ast_of (Const (c, _)) = Ast.Constant c
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   472
      | ast_of (Free (x, _)) = Ast.Variable x
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   473
      | ast_of (Var (xi, _)) = Ast.Variable (tune_var (Term.string_of_vname xi))
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   474
      | ast_of (t as _ $ _) =
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   475
          let val (f, args) = strip_comb t
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   476
          in Ast.mk_appl (ast_of f) (map ast_of args) end
42408
282b7a3065d3 explicit markup for loose bounds;
wenzelm
parents: 42382
diff changeset
   477
      | ast_of (Bound i) = Ast.Appl [Ast.Constant "_loose", Ast.Variable ("B." ^ string_of_int i)]
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   478
      | ast_of (Abs _) = raise Fail "simple_ast_of: Abs";
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   479
  in ast_of end;
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   480
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   481
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   482
(* sort_to_ast and typ_to_ast *)
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   483
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   484
fun ast_of_termT ctxt trf tm =
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   485
  let
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   486
    val ctxt' = Config.put show_sorts false ctxt;
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   487
    fun ast_of (t as Const ("_tfree", _) $ Free _) = simple_ast_of ctxt t
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   488
      | ast_of (t as Const ("_tvar", _) $ Var _) = simple_ast_of ctxt t
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   489
      | ast_of (Const (a, _)) = trans a []
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   490
      | ast_of (t as _ $ _) =
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   491
          (case strip_comb t of
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   492
            (Const (a, _), args) => trans a args
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   493
          | (f, args) => Ast.Appl (map ast_of (f :: args)))
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   494
      | ast_of t = simple_ast_of ctxt t
42254
f427c9890c46 more abstract print translation;
wenzelm
parents: 42253
diff changeset
   495
    and trans a args = ast_of (trf a ctxt' dummyT args)
f427c9890c46 more abstract print translation;
wenzelm
parents: 42253
diff changeset
   496
      handle Match => Ast.mk_appl (Ast.Constant a) (map ast_of args);
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   497
  in ast_of tm end;
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   498
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   499
fun sort_to_ast ctxt trf S = ast_of_termT ctxt trf (term_of_sort S);
42247
12fe41a92cd5 typed_print_translation: discontinued show_sorts argument;
wenzelm
parents: 42245
diff changeset
   500
fun typ_to_ast ctxt trf T = ast_of_termT ctxt trf (term_of_typ ctxt T);
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   501
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   502
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   503
(* term_to_ast *)
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   504
42252
wenzelm
parents: 42251
diff changeset
   505
fun term_to_ast idents is_syntax_const ctxt trf tm =
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   506
  let
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   507
    val show_types =
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   508
      Config.get ctxt show_types orelse Config.get ctxt show_sorts orelse
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   509
      Config.get ctxt show_all_types;
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   510
    val show_structs = Config.get ctxt show_structs;
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   511
    val show_free_types = Config.get ctxt show_free_types;
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   512
    val show_all_types = Config.get ctxt show_all_types;
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   513
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   514
    val {structs, fixes} = idents;
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   515
42249
12a073670584 simplified standard parse/unparse;
wenzelm
parents: 42248
diff changeset
   516
    fun mark_atoms ((t as Const (c, _)) $ u) =
42294
0f4372a2d2e4 simplified Pure syntax bootstrap;
wenzelm
parents: 42290
diff changeset
   517
          if member (op =) Syntax.token_markers c
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   518
          then t $ u else mark_atoms t $ mark_atoms u
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   519
      | mark_atoms (t $ u) = mark_atoms t $ mark_atoms u
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   520
      | mark_atoms (Abs (x, T, t)) = Abs (x, T, mark_atoms t)
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   521
      | mark_atoms (t as Const (c, T)) =
42252
wenzelm
parents: 42251
diff changeset
   522
          if is_syntax_const c then t
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   523
          else Const (Lexicon.mark_const c, T)
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   524
      | mark_atoms (t as Free (x, T)) =
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   525
          let val i = find_index (fn s => s = x) structs + 1 in
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   526
            if i = 0 andalso member (op =) fixes x then
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   527
              Const (Lexicon.mark_fixed x, T)
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   528
            else if i = 1 andalso not show_structs then
42476
d0bc1268ef09 clarified auxiliary structure Lexicon.Syntax;
wenzelm
parents: 42470
diff changeset
   529
              Syntax.const "_struct" $ Syntax.const "_indexdefault"
d0bc1268ef09 clarified auxiliary structure Lexicon.Syntax;
wenzelm
parents: 42470
diff changeset
   530
            else Syntax.const "_free" $ t
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   531
          end
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   532
      | mark_atoms (t as Var (xi, T)) =
42288
2074b31650e6 discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
wenzelm
parents: 42284
diff changeset
   533
          if xi = Syntax_Ext.dddot_indexname then Const ("_DDDOT", T)
42476
d0bc1268ef09 clarified auxiliary structure Lexicon.Syntax;
wenzelm
parents: 42470
diff changeset
   534
          else Syntax.const "_var" $ t
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   535
      | mark_atoms a = a;
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   536
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   537
    fun prune_typs (t_seen as (Const _, _)) = t_seen
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   538
      | prune_typs (t as Free (x, ty), seen) =
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   539
          if ty = dummyT then (t, seen)
42290
b1f544c84040 discontinued special treatment of structure Lexicon;
wenzelm
parents: 42289
diff changeset
   540
          else if not show_free_types orelse member (op aconv) seen t then (Syntax.free x, seen)
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   541
          else (t, t :: seen)
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   542
      | prune_typs (t as Var (xi, ty), seen) =
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   543
          if ty = dummyT then (t, seen)
42290
b1f544c84040 discontinued special treatment of structure Lexicon;
wenzelm
parents: 42289
diff changeset
   544
          else if not show_free_types orelse member (op aconv) seen t then (Syntax.var xi, seen)
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   545
          else (t, t :: seen)
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   546
      | prune_typs (t_seen as (Bound _, _)) = t_seen
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   547
      | prune_typs (Abs (x, ty, t), seen) =
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   548
          let val (t', seen') = prune_typs (t, seen);
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   549
          in (Abs (x, ty, t'), seen') end
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   550
      | prune_typs (t1 $ t2, seen) =
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   551
          let
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   552
            val (t1', seen') = prune_typs (t1, seen);
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   553
            val (t2', seen'') = prune_typs (t2, seen');
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   554
          in (t1' $ t2', seen'') end;
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   555
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   556
    fun ast_of tm =
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   557
      (case strip_comb tm of
42284
326f57825e1a explicit structure Syntax_Trans;
wenzelm
parents: 42282
diff changeset
   558
        (t as Abs _, ts) => Ast.mk_appl (ast_of (Syntax_Trans.abs_tr' ctxt t)) (map ast_of ts)
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   559
      | ((c as Const ("_free", _)), Free (x, T) :: ts) =>
42290
b1f544c84040 discontinued special treatment of structure Lexicon;
wenzelm
parents: 42289
diff changeset
   560
          Ast.mk_appl (constrain (c $ Syntax.free x) T) (map ast_of ts)
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   561
      | ((c as Const ("_var", _)), Var (xi, T) :: ts) =>
42290
b1f544c84040 discontinued special treatment of structure Lexicon;
wenzelm
parents: 42289
diff changeset
   562
          Ast.mk_appl (constrain (c $ Syntax.var xi) T) (map ast_of ts)
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   563
      | ((c as Const ("_bound", _)), Free (x, T) :: ts) =>
42290
b1f544c84040 discontinued special treatment of structure Lexicon;
wenzelm
parents: 42289
diff changeset
   564
          Ast.mk_appl (constrain (c $ Syntax.free x) T) (map ast_of ts)
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   565
      | (Const ("_idtdummy", T), ts) =>
42476
d0bc1268ef09 clarified auxiliary structure Lexicon.Syntax;
wenzelm
parents: 42470
diff changeset
   566
          Ast.mk_appl (constrain (Syntax.const "_idtdummy") T) (map ast_of ts)
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   567
      | (const as Const (c, T), ts) =>
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   568
          if show_all_types
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   569
          then Ast.mk_appl (constrain const T) (map ast_of ts)
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   570
          else trans c T ts
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   571
      | (t, ts) => Ast.mk_appl (simple_ast_of ctxt t) (map ast_of ts))
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   572
42254
f427c9890c46 more abstract print translation;
wenzelm
parents: 42253
diff changeset
   573
    and trans a T args = ast_of (trf a ctxt T args)
f427c9890c46 more abstract print translation;
wenzelm
parents: 42253
diff changeset
   574
      handle Match => Ast.mk_appl (Ast.Constant a) (map ast_of args)
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   575
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   576
    and constrain t T =
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   577
      if show_types andalso T <> dummyT then
42248
04bffad68aa4 discontinued old-style Syntax.constrainC;
wenzelm
parents: 42247
diff changeset
   578
        Ast.Appl [Ast.Constant "_constrain", simple_ast_of ctxt t,
42247
12fe41a92cd5 typed_print_translation: discontinued show_sorts argument;
wenzelm
parents: 42245
diff changeset
   579
          ast_of_termT ctxt trf (term_of_typ ctxt T)]
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   580
      else simple_ast_of ctxt t;
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   581
  in
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   582
    tm
42284
326f57825e1a explicit structure Syntax_Trans;
wenzelm
parents: 42282
diff changeset
   583
    |> Syntax_Trans.prop_tr'
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   584
    |> show_types ? (#1 o prune_typs o rpair [])
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   585
    |> mark_atoms
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   586
    |> ast_of
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   587
  end;
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   588
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   589
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   590
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   591
(** unparse **)
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   592
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   593
local
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   594
42267
9566078ad905 simplified printer context: uniform externing and static token translations;
wenzelm
parents: 42264
diff changeset
   595
fun free_or_skolem ctxt x =
9566078ad905 simplified printer context: uniform externing and static token translations;
wenzelm
parents: 42264
diff changeset
   596
  let
9566078ad905 simplified printer context: uniform externing and static token translations;
wenzelm
parents: 42264
diff changeset
   597
    val m =
9566078ad905 simplified printer context: uniform externing and static token translations;
wenzelm
parents: 42264
diff changeset
   598
      if Variable.is_fixed ctxt x orelse Syntax.is_pretty_global ctxt
9566078ad905 simplified printer context: uniform externing and static token translations;
wenzelm
parents: 42264
diff changeset
   599
      then Markup.fixed x
9566078ad905 simplified printer context: uniform externing and static token translations;
wenzelm
parents: 42264
diff changeset
   600
      else Markup.hilite;
9566078ad905 simplified printer context: uniform externing and static token translations;
wenzelm
parents: 42264
diff changeset
   601
  in
9566078ad905 simplified printer context: uniform externing and static token translations;
wenzelm
parents: 42264
diff changeset
   602
    if can Name.dest_skolem x
42488
4638622bcaa1 reorganized fixes as specialized (global) name space;
wenzelm
parents: 42476
diff changeset
   603
    then ([m, Markup.skolem], Variable.revert_fixed ctxt x)
42267
9566078ad905 simplified printer context: uniform externing and static token translations;
wenzelm
parents: 42264
diff changeset
   604
    else ([m, Markup.free], x)
9566078ad905 simplified printer context: uniform externing and static token translations;
wenzelm
parents: 42264
diff changeset
   605
  end;
9566078ad905 simplified printer context: uniform externing and static token translations;
wenzelm
parents: 42264
diff changeset
   606
9566078ad905 simplified printer context: uniform externing and static token translations;
wenzelm
parents: 42264
diff changeset
   607
fun var_or_skolem s =
9566078ad905 simplified printer context: uniform externing and static token translations;
wenzelm
parents: 42264
diff changeset
   608
  (case Lexicon.read_variable s of
9566078ad905 simplified printer context: uniform externing and static token translations;
wenzelm
parents: 42264
diff changeset
   609
    SOME (x, i) =>
9566078ad905 simplified printer context: uniform externing and static token translations;
wenzelm
parents: 42264
diff changeset
   610
      (case try Name.dest_skolem x of
9566078ad905 simplified printer context: uniform externing and static token translations;
wenzelm
parents: 42264
diff changeset
   611
        NONE => (Markup.var, s)
9566078ad905 simplified printer context: uniform externing and static token translations;
wenzelm
parents: 42264
diff changeset
   612
      | SOME x' => (Markup.skolem, Term.string_of_vname (x', i)))
9566078ad905 simplified printer context: uniform externing and static token translations;
wenzelm
parents: 42264
diff changeset
   613
  | NONE => (Markup.var, s));
9566078ad905 simplified printer context: uniform externing and static token translations;
wenzelm
parents: 42264
diff changeset
   614
9566078ad905 simplified printer context: uniform externing and static token translations;
wenzelm
parents: 42264
diff changeset
   615
fun unparse_t t_to_ast prt_t markup ctxt t =
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   616
  let
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 42359
diff changeset
   617
    val syn = Proof_Context.syn_of ctxt;
42267
9566078ad905 simplified printer context: uniform externing and static token translations;
wenzelm
parents: 42264
diff changeset
   618
9566078ad905 simplified printer context: uniform externing and static token translations;
wenzelm
parents: 42264
diff changeset
   619
    fun token_trans "_tfree" x = SOME (Pretty.mark_str (Markup.tfree, x))
9566078ad905 simplified printer context: uniform externing and static token translations;
wenzelm
parents: 42264
diff changeset
   620
      | token_trans "_tvar" x = SOME (Pretty.mark_str (Markup.tvar, x))
9566078ad905 simplified printer context: uniform externing and static token translations;
wenzelm
parents: 42264
diff changeset
   621
      | 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: 42264
diff changeset
   622
      | token_trans "_bound" x = SOME (Pretty.mark_str (Markup.bound, x))
42408
282b7a3065d3 explicit markup for loose bounds;
wenzelm
parents: 42382
diff changeset
   623
      | token_trans "_loose" x = SOME (Pretty.mark_str (Markup.malformed, x))
42267
9566078ad905 simplified printer context: uniform externing and static token translations;
wenzelm
parents: 42264
diff changeset
   624
      | token_trans "_var" x = SOME (Pretty.mark_str (var_or_skolem x))
9566078ad905 simplified printer context: uniform externing and static token translations;
wenzelm
parents: 42264
diff changeset
   625
      | token_trans "_numeral" x = SOME (Pretty.mark_str (Markup.numeral, x))
9566078ad905 simplified printer context: uniform externing and static token translations;
wenzelm
parents: 42264
diff changeset
   626
      | 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: 42264
diff changeset
   627
      | token_trans _ _ = NONE;
9566078ad905 simplified printer context: uniform externing and static token translations;
wenzelm
parents: 42264
diff changeset
   628
42301
d7ca0c03d4ea unparse: more accurate markup for syntax consts, notably binders;
wenzelm
parents: 42298
diff changeset
   629
    fun markup_extern c =
d7ca0c03d4ea unparse: more accurate markup for syntax consts, notably binders;
wenzelm
parents: 42298
diff changeset
   630
      (case Syntax.lookup_const syn c of
d7ca0c03d4ea unparse: more accurate markup for syntax consts, notably binders;
wenzelm
parents: 42298
diff changeset
   631
        SOME "" => ([], c)
d7ca0c03d4ea unparse: more accurate markup for syntax consts, notably binders;
wenzelm
parents: 42298
diff changeset
   632
      | SOME b => markup_extern b
d7ca0c03d4ea unparse: more accurate markup for syntax consts, notably binders;
wenzelm
parents: 42298
diff changeset
   633
      | NONE => c |> Lexicon.unmark
43548
f231a7594e54 type classes: entity markup instead of old-style token markup;
wenzelm
parents: 42493
diff changeset
   634
         {case_class = fn x => (markup_class ctxt x, Proof_Context.extern_class ctxt x),
43552
156c822f181a entity markup for "type", "constant";
wenzelm
parents: 43548
diff changeset
   635
          case_type = fn x => (markup_type ctxt x, Proof_Context.extern_type ctxt x),
156c822f181a entity markup for "type", "constant";
wenzelm
parents: 43548
diff changeset
   636
          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: 42298
diff changeset
   637
          case_fixed = fn x => free_or_skolem ctxt x,
d7ca0c03d4ea unparse: more accurate markup for syntax consts, notably binders;
wenzelm
parents: 42298
diff changeset
   638
          case_default = fn x => ([], x)});
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   639
  in
42255
097c93dcd877 eliminated slightly odd Syntax.rep_syntax;
wenzelm
parents: 42254
diff changeset
   640
    t_to_ast ctxt (Syntax.print_translation syn) t
097c93dcd877 eliminated slightly odd Syntax.rep_syntax;
wenzelm
parents: 42254
diff changeset
   641
    |> Ast.normalize ctxt (Syntax.print_rules syn)
42267
9566078ad905 simplified printer context: uniform externing and static token translations;
wenzelm
parents: 42264
diff changeset
   642
    |> prt_t ctxt (Syntax.prtabs syn) (Syntax.print_ast_translation syn) token_trans markup_extern
42253
c539d3c9c023 more abstract syntax translation;
wenzelm
parents: 42252
diff changeset
   643
    |> Pretty.markup markup
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   644
  end;
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   645
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   646
in
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   647
42267
9566078ad905 simplified printer context: uniform externing and static token translations;
wenzelm
parents: 42264
diff changeset
   648
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: 42264
diff changeset
   649
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
   650
dd8029f71e1c separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff changeset
   651
fun unparse_term ctxt =
dd8029f71e1c separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff changeset
   652
  let
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 42359
diff changeset
   653
    val thy = Proof_Context.theory_of ctxt;
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 42359
diff changeset
   654
    val syn = Proof_Context.syn_of ctxt;
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 42359
diff changeset
   655
    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
   656
  in
42298
d622145603ee more accurate markup for syntax consts, notably binders which point back to the original logical entity;
wenzelm
parents: 42296
diff changeset
   657
    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: 42264
diff changeset
   658
      (Printer.pretty_term_ast (not (Pure_Thy.old_appl_syntax thy)))
9566078ad905 simplified printer context: uniform externing and static token translations;
wenzelm
parents: 42264
diff changeset
   659
      Markup.term ctxt
42241
dd8029f71e1c separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff changeset
   660
  end;
dd8029f71e1c separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff changeset
   661
42249
12a073670584 simplified standard parse/unparse;
wenzelm
parents: 42248
diff changeset
   662
end;
12a073670584 simplified standard parse/unparse;
wenzelm
parents: 42248
diff changeset
   663
42241
dd8029f71e1c separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff changeset
   664
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   665
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   666
(** translations **)
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   667
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   668
(* type propositions *)
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   669
42247
12fe41a92cd5 typed_print_translation: discontinued show_sorts argument;
wenzelm
parents: 42245
diff changeset
   670
fun type_prop_tr' ctxt T [Const ("\\<^const>Pure.sort_constraint", _)] =
42476
d0bc1268ef09 clarified auxiliary structure Lexicon.Syntax;
wenzelm
parents: 42470
diff changeset
   671
      Syntax.const "_sort_constraint" $ term_of_typ (Config.put show_sorts true ctxt) T
42247
12fe41a92cd5 typed_print_translation: discontinued show_sorts argument;
wenzelm
parents: 42245
diff changeset
   672
  | type_prop_tr' ctxt T [t] =
42476
d0bc1268ef09 clarified auxiliary structure Lexicon.Syntax;
wenzelm
parents: 42470
diff changeset
   673
      Syntax.const "_ofclass" $ term_of_typ ctxt T $ t
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   674
  | type_prop_tr' _ T ts = raise TYPE ("type_prop_tr'", [T], ts);
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   675
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   676
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   677
(* type reflection *)
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   678
42247
12fe41a92cd5 typed_print_translation: discontinued show_sorts argument;
wenzelm
parents: 42245
diff changeset
   679
fun type_tr' ctxt (Type ("itself", [T])) ts =
42476
d0bc1268ef09 clarified auxiliary structure Lexicon.Syntax;
wenzelm
parents: 42470
diff changeset
   680
      Term.list_comb (Syntax.const "_TYPE" $ term_of_typ ctxt T, ts)
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   681
  | type_tr' _ _ _ = raise Match;
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   682
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   683
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   684
(* type constraints *)
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   685
42247
12fe41a92cd5 typed_print_translation: discontinued show_sorts argument;
wenzelm
parents: 42245
diff changeset
   686
fun type_constraint_tr' ctxt (Type ("fun", [T, _])) (t :: ts) =
42476
d0bc1268ef09 clarified auxiliary structure Lexicon.Syntax;
wenzelm
parents: 42470
diff changeset
   687
      Term.list_comb (Syntax.const "_constrain" $ t $ term_of_typ ctxt T, ts)
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   688
  | type_constraint_tr' _ _ _ = raise Match;
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   689
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   690
42295
8fdbb3b10beb moved CONST syntax/translations to their proper place;
wenzelm
parents: 42294
diff changeset
   691
(* authentic syntax *)
8fdbb3b10beb moved CONST syntax/translations to their proper place;
wenzelm
parents: 42294
diff changeset
   692
8fdbb3b10beb moved CONST syntax/translations to their proper place;
wenzelm
parents: 42294
diff changeset
   693
fun const_ast_tr intern ctxt [Ast.Variable c] =
8fdbb3b10beb moved CONST syntax/translations to their proper place;
wenzelm
parents: 42294
diff changeset
   694
      let
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 42359
diff changeset
   695
        val Const (c', _) = Proof_Context.read_const_proper ctxt false c;
42295
8fdbb3b10beb moved CONST syntax/translations to their proper place;
wenzelm
parents: 42294
diff changeset
   696
        val d = if intern then Lexicon.mark_const c' else c;
8fdbb3b10beb moved CONST syntax/translations to their proper place;
wenzelm
parents: 42294
diff changeset
   697
      in Ast.Constant d end
42470
cc78b0ed0fad more precise error positions;
wenzelm
parents: 42410
diff changeset
   698
  | const_ast_tr intern ctxt [Ast.Appl [Ast.Constant "_constrain", x, T as Ast.Variable pos]] =
cc78b0ed0fad more precise error positions;
wenzelm
parents: 42410
diff changeset
   699
      (Ast.Appl [Ast.Constant "_constrain", const_ast_tr intern ctxt [x], T]
cc78b0ed0fad more precise error positions;
wenzelm
parents: 42410
diff changeset
   700
        handle ERROR msg =>
cc78b0ed0fad more precise error positions;
wenzelm
parents: 42410
diff changeset
   701
          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: 42294
diff changeset
   702
  | const_ast_tr _ _ asts = raise Ast.AST ("const_ast_tr", asts);
8fdbb3b10beb moved CONST syntax/translations to their proper place;
wenzelm
parents: 42294
diff changeset
   703
8fdbb3b10beb moved CONST syntax/translations to their proper place;
wenzelm
parents: 42294
diff changeset
   704
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   705
(* setup translations *)
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   706
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   707
val _ = Context.>> (Context.map_theory
42295
8fdbb3b10beb moved CONST syntax/translations to their proper place;
wenzelm
parents: 42294
diff changeset
   708
 (Sign.add_advanced_trfuns
8fdbb3b10beb moved CONST syntax/translations to their proper place;
wenzelm
parents: 42294
diff changeset
   709
  ([("_context_const", const_ast_tr true),
8fdbb3b10beb moved CONST syntax/translations to their proper place;
wenzelm
parents: 42294
diff changeset
   710
    ("_context_xconst", const_ast_tr false)], [], [], []) #>
8fdbb3b10beb moved CONST syntax/translations to their proper place;
wenzelm
parents: 42294
diff changeset
   711
  Sign.add_advanced_trfunsT
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   712
   [("_type_prop", type_prop_tr'),
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   713
    ("\\<^const>TYPE", type_tr'),
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   714
    ("_type_constraint_", type_constraint_tr')]));
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   715
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   716
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   717
42382
dcd983ee2c29 provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents: 42379
diff changeset
   718
(** check/uncheck **)
dcd983ee2c29 provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents: 42379
diff changeset
   719
45429
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   720
(* context-sensitive (un)checking *)
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   721
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   722
type key = int * bool;
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   723
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   724
structure Checks = Generic_Data
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   725
(
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   726
  type 'a check = 'a list -> Proof.context -> ('a list * Proof.context) option;
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   727
  type T =
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   728
    ((key * ((string * typ check) * stamp) list) list *
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   729
     (key * ((string * term check) * stamp) list) list);
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   730
  val empty = ([], []);
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   731
  val extend = I;
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   732
  fun merge ((typ_checks1, term_checks1), (typ_checks2, term_checks2)) : T =
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   733
    (AList.join (op =) (K (Library.merge (eq_snd (op =)))) (typ_checks1, typ_checks2),
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   734
     AList.join (op =) (K (Library.merge (eq_snd (op =)))) (term_checks1, term_checks2));
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   735
);
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   736
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   737
fun print_checks ctxt =
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   738
  let
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   739
    fun split_checks checks =
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   740
      List.partition (fn ((_, un), _) => not un) checks
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   741
      |> pairself (map (fn ((i, _), fs) => (i, map (fst o fst) fs))
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   742
          #> sort (int_ord o pairself fst));
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   743
    fun pretty_checks kind checks =
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   744
      checks |> map (fn (i, names) => Pretty.block
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   745
        [Pretty.str (kind ^ " (stage " ^ signed_string_of_int i ^ "):"),
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   746
          Pretty.brk 1, Pretty.strs names]);
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   747
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   748
    val (typs, terms) = Checks.get (Context.Proof ctxt);
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   749
    val (typ_checks, typ_unchecks) = split_checks typs;
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   750
    val (term_checks, term_unchecks) = split_checks terms;
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   751
  in
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   752
    pretty_checks "typ_checks" typ_checks @
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   753
    pretty_checks "term_checks" term_checks @
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   754
    pretty_checks "typ_unchecks" typ_unchecks @
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   755
    pretty_checks "term_unchecks" term_unchecks
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   756
  end |> Pretty.chunks |> Pretty.writeln;
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   757
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   758
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   759
local
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   760
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   761
fun context_check which (key: key) name f =
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   762
  Checks.map (which (AList.map_default op = (key, []) (cons ((name, f), stamp ()))));
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   763
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   764
fun simple_check eq f xs ctxt =
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   765
  let val xs' = f ctxt xs
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   766
  in if eq_list eq (xs, xs') then NONE else SOME (xs', ctxt) end;
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   767
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   768
in
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   769
45444
ac069060e08a tuned signature;
wenzelm
parents: 45429
diff changeset
   770
fun typ_check' stage = context_check apfst (stage, false);
ac069060e08a tuned signature;
wenzelm
parents: 45429
diff changeset
   771
fun term_check' stage = context_check apsnd (stage, false);
ac069060e08a tuned signature;
wenzelm
parents: 45429
diff changeset
   772
fun typ_uncheck' stage = context_check apfst (stage, true);
ac069060e08a tuned signature;
wenzelm
parents: 45429
diff changeset
   773
fun term_uncheck' stage = context_check apsnd (stage, true);
45429
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   774
45444
ac069060e08a tuned signature;
wenzelm
parents: 45429
diff changeset
   775
fun typ_check key name f = typ_check' key name (simple_check (op =) f);
ac069060e08a tuned signature;
wenzelm
parents: 45429
diff changeset
   776
fun term_check key name f = term_check' key name (simple_check (op aconv) f);
ac069060e08a tuned signature;
wenzelm
parents: 45429
diff changeset
   777
fun typ_uncheck key name f = typ_uncheck' key name (simple_check (op =) f);
ac069060e08a tuned signature;
wenzelm
parents: 45429
diff changeset
   778
fun term_uncheck key name f = term_uncheck' key name (simple_check (op aconv) f);
45429
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   779
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   780
end;
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   781
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   782
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   783
local
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   784
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   785
fun check_stage fs = perhaps_loop (perhaps_apply (map uncurry fs));
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   786
fun check_all fs = perhaps_apply (map check_stage fs);
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   787
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   788
fun check which uncheck ctxt0 xs0 =
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   789
  let
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   790
    val funs = which (Checks.get (Context.Proof ctxt0))
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   791
      |> map_filter (fn ((i, u), fs) => if uncheck = u then SOME (i, map (snd o fst) fs) else NONE)
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   792
      |> Library.sort (int_ord o pairself fst) |> map snd
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   793
      |> not uncheck ? map rev;
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   794
  in #1 (perhaps (check_all funs) (xs0, ctxt0)) end;
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   795
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   796
val apply_typ_check = check fst false;
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   797
val apply_term_check = check snd false;
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   798
val apply_typ_uncheck = check fst true;
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   799
val apply_term_uncheck = check snd true;
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   800
45445
41e641a870de pass term positions into check phase, where resulting types are reported accordingly, and eventually shown as tooltips;
wenzelm
parents: 45444
diff changeset
   801
fun prepare_sorts ctxt tys =
45427
fca432074fb2 sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents: 45423
diff changeset
   802
  let
fca432074fb2 sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents: 45423
diff changeset
   803
    fun constraint (xi: indexname, S) = S <> dummyS ? insert (op =) (xi, S);
fca432074fb2 sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents: 45423
diff changeset
   804
    val env =
fca432074fb2 sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents: 45423
diff changeset
   805
      (fold o fold_atyps)
fca432074fb2 sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents: 45423
diff changeset
   806
        (fn TFree (x, S) => constraint ((x, ~1), S)
fca432074fb2 sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents: 45423
diff changeset
   807
          | TVar v => constraint v
fca432074fb2 sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents: 45423
diff changeset
   808
          | _ => I) tys [];
fca432074fb2 sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents: 45423
diff changeset
   809
    val get_sort = Proof_Context.get_sort ctxt env;
fca432074fb2 sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents: 45423
diff changeset
   810
  in
fca432074fb2 sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents: 45423
diff changeset
   811
    (map o map_atyps)
fca432074fb2 sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents: 45423
diff changeset
   812
      (fn TFree (x, _) => TFree (x, get_sort (x, ~1))
fca432074fb2 sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents: 45423
diff changeset
   813
        | TVar (xi, _) => TVar (xi, get_sort xi)
fca432074fb2 sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents: 45423
diff changeset
   814
        | T => T) tys
fca432074fb2 sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents: 45423
diff changeset
   815
  end;
fca432074fb2 sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents: 45423
diff changeset
   816
45429
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   817
in
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   818
45423
92f91f990165 tuned signature -- emphasize internal role of these operations;
wenzelm
parents: 45412
diff changeset
   819
fun check_typs ctxt =
45445
41e641a870de pass term positions into check phase, where resulting types are reported accordingly, and eventually shown as tooltips;
wenzelm
parents: 45444
diff changeset
   820
  prepare_sorts ctxt #>
45429
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   821
  apply_typ_check ctxt #>
45427
fca432074fb2 sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents: 45423
diff changeset
   822
  Term_Sharing.typs (Proof_Context.theory_of ctxt);
45423
92f91f990165 tuned signature -- emphasize internal role of these operations;
wenzelm
parents: 45412
diff changeset
   823
45445
41e641a870de pass term positions into check phase, where resulting types are reported accordingly, and eventually shown as tooltips;
wenzelm
parents: 45444
diff changeset
   824
fun check_terms ctxt raw_ts =
41e641a870de pass term positions into check phase, where resulting types are reported accordingly, and eventually shown as tooltips;
wenzelm
parents: 45444
diff changeset
   825
  let
41e641a870de pass term positions into check phase, where resulting types are reported accordingly, and eventually shown as tooltips;
wenzelm
parents: 45444
diff changeset
   826
    val (ts, ps) = Type_Infer_Context.prepare_positions ctxt raw_ts;
41e641a870de pass term positions into check phase, where resulting types are reported accordingly, and eventually shown as tooltips;
wenzelm
parents: 45444
diff changeset
   827
    val tys = map (Logic.mk_type o snd) ps;
41e641a870de pass term positions into check phase, where resulting types are reported accordingly, and eventually shown as tooltips;
wenzelm
parents: 45444
diff changeset
   828
    val (ts', tys') =
41e641a870de pass term positions into check phase, where resulting types are reported accordingly, and eventually shown as tooltips;
wenzelm
parents: 45444
diff changeset
   829
      Term.burrow_types (prepare_sorts ctxt) ts @ tys
41e641a870de pass term positions into check phase, where resulting types are reported accordingly, and eventually shown as tooltips;
wenzelm
parents: 45444
diff changeset
   830
      |> apply_term_check ctxt
41e641a870de pass term positions into check phase, where resulting types are reported accordingly, and eventually shown as tooltips;
wenzelm
parents: 45444
diff changeset
   831
      |> chop (length ts);
41e641a870de pass term positions into check phase, where resulting types are reported accordingly, and eventually shown as tooltips;
wenzelm
parents: 45444
diff changeset
   832
    val _ =
41e641a870de pass term positions into check phase, where resulting types are reported accordingly, and eventually shown as tooltips;
wenzelm
parents: 45444
diff changeset
   833
      map2 (fn (pos, _) => fn ty =>
41e641a870de pass term positions into check phase, where resulting types are reported accordingly, and eventually shown as tooltips;
wenzelm
parents: 45444
diff changeset
   834
        if Position.is_reported pos then
41e641a870de pass term positions into check phase, where resulting types are reported accordingly, and eventually shown as tooltips;
wenzelm
parents: 45444
diff changeset
   835
          Markup.markup (Position.markup pos Markup.typing)
41e641a870de pass term positions into check phase, where resulting types are reported accordingly, and eventually shown as tooltips;
wenzelm
parents: 45444
diff changeset
   836
            (Syntax.string_of_typ ctxt (Logic.dest_type ty))
41e641a870de pass term positions into check phase, where resulting types are reported accordingly, and eventually shown as tooltips;
wenzelm
parents: 45444
diff changeset
   837
        else "") ps tys'
41e641a870de pass term positions into check phase, where resulting types are reported accordingly, and eventually shown as tooltips;
wenzelm
parents: 45444
diff changeset
   838
      |> implode |> Output.report
41e641a870de pass term positions into check phase, where resulting types are reported accordingly, and eventually shown as tooltips;
wenzelm
parents: 45444
diff changeset
   839
  in Term_Sharing.terms (Proof_Context.theory_of ctxt) ts' end;
45423
92f91f990165 tuned signature -- emphasize internal role of these operations;
wenzelm
parents: 45412
diff changeset
   840
42382
dcd983ee2c29 provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents: 42379
diff changeset
   841
fun check_props ctxt = map (Type.constraint propT) #> check_terms ctxt;
dcd983ee2c29 provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents: 42379
diff changeset
   842
45429
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   843
val uncheck_typs = apply_typ_uncheck;
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   844
val uncheck_terms = apply_term_uncheck;
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   845
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   846
end;
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   847
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   848
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   849
(* standard phases *)
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   850
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   851
val _ = Context.>>
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   852
 (typ_check 0 "standard" Proof_Context.standard_typ_check #>
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   853
  term_check 0 "standard"
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   854
    (fn ctxt => Type_Infer_Context.infer_types ctxt #> map (Proof_Context.expand_abbrevs ctxt)) #>
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   855
  term_check 100 "standard_finish" Proof_Context.standard_term_check_finish #>
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   856
  term_uncheck 0 "standard" Proof_Context.standard_term_uncheck);
42382
dcd983ee2c29 provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents: 42379
diff changeset
   857
dcd983ee2c29 provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents: 42379
diff changeset
   858
dcd983ee2c29 provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents: 42379
diff changeset
   859
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   860
(** install operations **)
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   861
42241
dd8029f71e1c separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff changeset
   862
val _ = Syntax.install_operations
dd8029f71e1c separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff changeset
   863
  {parse_sort = parse_sort,
dd8029f71e1c separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff changeset
   864
   parse_typ = parse_typ,
42281
69d4543811d0 simplified read_term vs. read_prop;
wenzelm
parents: 42267
diff changeset
   865
   parse_term = parse_term false,
69d4543811d0 simplified read_term vs. read_prop;
wenzelm
parents: 42267
diff changeset
   866
   parse_prop = parse_term true,
42241
dd8029f71e1c separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff changeset
   867
   unparse_sort = unparse_sort,
dd8029f71e1c separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff changeset
   868
   unparse_typ = unparse_typ,
42382
dcd983ee2c29 provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents: 42379
diff changeset
   869
   unparse_term = unparse_term,
dcd983ee2c29 provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents: 42379
diff changeset
   870
   check_typs = check_typs,
dcd983ee2c29 provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents: 42379
diff changeset
   871
   check_terms = check_terms,
dcd983ee2c29 provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents: 42379
diff changeset
   872
   check_props = check_props,
dcd983ee2c29 provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents: 42379
diff changeset
   873
   uncheck_typs = uncheck_typs,
dcd983ee2c29 provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents: 42379
diff changeset
   874
   uncheck_terms = uncheck_terms};
42241
dd8029f71e1c separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff changeset
   875
dd8029f71e1c separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff changeset
   876
end;