src/Pure/Syntax/syntax_phases.ML
author blanchet
Thu, 30 Jan 2014 15:01:40 +0100
changeset 55184 6e2295db4cf8
parent 55014 a93f496f6c30
child 55550 bcc643ac071a
permissions -rw-r--r--
keep formula right before skolemization, because the universal variables might be different (or differently ordered) as in the original axiom or negated conjecture from which it was skolemized
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 =
50201
c26369c9eda6 Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents: 49690
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
50201
c26369c9eda6 Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents: 49690
diff changeset
    59
fun markup_var xi = [Markup.name (Term.string_of_vname xi) Markup.var];
42298
d622145603ee more accurate markup for syntax consts, notably binders which point back to the original logical entity;
wenzelm
parents: 42296
diff changeset
    60
45412
7797f5351ec4 entity markup for bound variables;
wenzelm
parents: 44736
diff changeset
    61
fun markup_bound def ps (name, id) =
50201
c26369c9eda6 Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents: 49690
diff changeset
    62
  let val entity = Markup.entity Markup.boundN name in
c26369c9eda6 Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents: 49690
diff changeset
    63
    Markup.bound ::
45412
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", _)) = []
49674
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 49662
diff changeset
    95
      | sort (Const ("_sort", _) $ cs) = classes cs
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
    96
      | sort (Const (s, _)) = [class s]
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
49685
4341e4d86872 allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents: 49674
diff changeset
   103
fun decode_pos (Free (s, _)) =
4341e4d86872 allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents: 49674
diff changeset
   104
      if is_some (Term_Position.decode s) then SOME s else NONE
4341e4d86872 allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents: 49674
diff changeset
   105
  | decode_pos _ = NONE;
4341e4d86872 allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents: 49674
diff changeset
   106
45427
fca432074fb2 sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents: 45423
diff changeset
   107
fun decode_typ tm =
fca432074fb2 sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents: 45423
diff changeset
   108
  let
fca432074fb2 sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents: 45423
diff changeset
   109
    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
   110
49685
4341e4d86872 allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents: 49674
diff changeset
   111
    fun typ ps sort tm =
4341e4d86872 allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents: 49674
diff changeset
   112
      (case tm of
4341e4d86872 allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents: 49674
diff changeset
   113
        Const ("_tfree", _) $ t => typ ps sort t
4341e4d86872 allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents: 49674
diff changeset
   114
      | Const ("_tvar", _) $ t => typ ps sort t
4341e4d86872 allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents: 49674
diff changeset
   115
      | Const ("_ofsort", _) $ t $ s =>
4341e4d86872 allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents: 49674
diff changeset
   116
          (case decode_pos s of
4341e4d86872 allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents: 49674
diff changeset
   117
            SOME p => typ (p :: ps) sort t
4341e4d86872 allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents: 49674
diff changeset
   118
          | NONE =>
4341e4d86872 allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents: 49674
diff changeset
   119
              if is_none sort then typ ps (SOME (decode_sort s)) t
4341e4d86872 allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents: 49674
diff changeset
   120
              else err ())
4341e4d86872 allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents: 49674
diff changeset
   121
      | Const ("_dummy_ofsort", _) $ s => TFree ("'_dummy_", decode_sort s)
4341e4d86872 allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents: 49674
diff changeset
   122
      | Free (x, _) => TFree (x, ps @ the_default dummyS sort)
4341e4d86872 allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents: 49674
diff changeset
   123
      | Var (xi, _) => TVar (xi, ps @ the_default dummyS sort)
4341e4d86872 allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents: 49674
diff changeset
   124
      | _ =>
4341e4d86872 allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents: 49674
diff changeset
   125
          if null ps andalso is_none sort then
4341e4d86872 allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents: 49674
diff changeset
   126
            let
4341e4d86872 allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents: 49674
diff changeset
   127
              val (head, args) = Term.strip_comb tm;
4341e4d86872 allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents: 49674
diff changeset
   128
              val a =
4341e4d86872 allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents: 49674
diff changeset
   129
                (case head of
4341e4d86872 allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents: 49674
diff changeset
   130
                  Const (c, _) => (Lexicon.unmark_type c handle Fail _ => err ())
4341e4d86872 allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents: 49674
diff changeset
   131
                | _ => err ());
4341e4d86872 allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents: 49674
diff changeset
   132
            in Type (a, map (typ [] NONE) args) end
4341e4d86872 allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents: 49674
diff changeset
   133
          else err ());
4341e4d86872 allow position constraints to coexist with 0 or 1 sort constraints;
wenzelm
parents: 49674
diff changeset
   134
  in typ [] NONE tm end;
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
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   137
(* parsetree_to_ast *)
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   138
52188
2da0033370a0 report markup for ast translations;
wenzelm
parents: 52186
diff changeset
   139
fun parsetree_to_ast ctxt trf parsetree =
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   140
  let
44736
c2a3f1c84179 bulk reports for improved message throughput;
wenzelm
parents: 44735
diff changeset
   141
    val reports = Unsynchronized.ref ([]: Position.report list);
44735
66862d02678c tuned signature;
wenzelm
parents: 44564
diff changeset
   142
    fun report pos = Position.store_reports reports [pos];
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   143
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   144
    fun trans a args =
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   145
      (case trf a of
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   146
        NONE => Ast.mk_appl (Ast.Constant a) args
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   147
      | SOME f => f ctxt args);
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   148
49674
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 49662
diff changeset
   149
    fun asts_of_token tok =
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 49662
diff changeset
   150
      if Lexicon.valued_token tok
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 49662
diff changeset
   151
      then [Ast.Variable (Lexicon.str_of_token tok)]
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 49662
diff changeset
   152
      else [];
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 49662
diff changeset
   153
52162
896ebb4646d8 position constraint for dummy_pattern -- more PIDE markup;
wenzelm
parents: 52160
diff changeset
   154
    fun ast_of_position tok =
896ebb4646d8 position constraint for dummy_pattern -- more PIDE markup;
wenzelm
parents: 52160
diff changeset
   155
      Ast.Variable (Term_Position.encode (Lexicon.pos_of_token tok));
896ebb4646d8 position constraint for dummy_pattern -- more PIDE markup;
wenzelm
parents: 52160
diff changeset
   156
52163
72e83c82c1d4 position constraint for bound dummy -- more PIDE markup;
wenzelm
parents: 52162
diff changeset
   157
    fun ast_of_dummy a tok =
52188
2da0033370a0 report markup for ast translations;
wenzelm
parents: 52186
diff changeset
   158
      Ast.Appl [Ast.Constant "_constrain", Ast.Constant a, ast_of_position tok];
52163
72e83c82c1d4 position constraint for bound dummy -- more PIDE markup;
wenzelm
parents: 52162
diff changeset
   159
49674
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 49662
diff changeset
   160
    fun asts_of_position c tok =
52188
2da0033370a0 report markup for ast translations;
wenzelm
parents: 52186
diff changeset
   161
      [Ast.Appl [Ast.Constant c, ast_of (Parser.Tip tok), ast_of_position tok]]
49674
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 49662
diff changeset
   162
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 49662
diff changeset
   163
    and asts_of (Parser.Node ("_class_name", [Parser.Tip tok])) =
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   164
          let
42470
cc78b0ed0fad more precise error positions;
wenzelm
parents: 42410
diff changeset
   165
            val pos = Lexicon.pos_of_token tok;
cc78b0ed0fad more precise error positions;
wenzelm
parents: 42410
diff changeset
   166
            val c = Proof_Context.read_class ctxt (Lexicon.str_of_token tok)
48992
0518bf89c777 renamed Position.str_of to Position.here;
wenzelm
parents: 48768
diff changeset
   167
              handle ERROR msg => error (msg ^ Position.here pos);
42470
cc78b0ed0fad more precise error positions;
wenzelm
parents: 42410
diff changeset
   168
            val _ = report pos (markup_class ctxt) c;
42282
4fa41e068ff0 report literal tokens according to parsetree head;
wenzelm
parents: 42281
diff changeset
   169
          in [Ast.Constant (Lexicon.mark_class c)] end
4fa41e068ff0 report literal tokens according to parsetree head;
wenzelm
parents: 42281
diff changeset
   170
      | asts_of (Parser.Node ("_type_name", [Parser.Tip tok])) =
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   171
          let
42470
cc78b0ed0fad more precise error positions;
wenzelm
parents: 42410
diff changeset
   172
            val pos = Lexicon.pos_of_token tok;
cc78b0ed0fad more precise error positions;
wenzelm
parents: 42410
diff changeset
   173
            val Type (c, _) =
cc78b0ed0fad more precise error positions;
wenzelm
parents: 42410
diff changeset
   174
              Proof_Context.read_type_name_proper ctxt false (Lexicon.str_of_token tok)
48992
0518bf89c777 renamed Position.str_of to Position.here;
wenzelm
parents: 48768
diff changeset
   175
                handle ERROR msg => error (msg ^ Position.here pos);
42470
cc78b0ed0fad more precise error positions;
wenzelm
parents: 42410
diff changeset
   176
            val _ = report pos (markup_type ctxt) c;
42282
4fa41e068ff0 report literal tokens according to parsetree head;
wenzelm
parents: 42281
diff changeset
   177
          in [Ast.Constant (Lexicon.mark_type c)] end
49674
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 49662
diff changeset
   178
      | asts_of (Parser.Node ("_position", [Parser.Tip tok])) = asts_of_position "_constrain" tok
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 49662
diff changeset
   179
      | asts_of (Parser.Node ("_position_sort", [Parser.Tip tok])) = asts_of_position "_ofsort" tok
52162
896ebb4646d8 position constraint for dummy_pattern -- more PIDE markup;
wenzelm
parents: 52160
diff changeset
   180
      | asts_of (Parser.Node (a as "\\<^const>dummy_pattern", [Parser.Tip tok])) =
52163
72e83c82c1d4 position constraint for bound dummy -- more PIDE markup;
wenzelm
parents: 52162
diff changeset
   181
          [ast_of_dummy a tok]
72e83c82c1d4 position constraint for bound dummy -- more PIDE markup;
wenzelm
parents: 52162
diff changeset
   182
      | asts_of (Parser.Node (a as "_idtdummy", [Parser.Tip tok])) =
72e83c82c1d4 position constraint for bound dummy -- more PIDE markup;
wenzelm
parents: 52162
diff changeset
   183
          [ast_of_dummy a tok]
72e83c82c1d4 position constraint for bound dummy -- more PIDE markup;
wenzelm
parents: 52162
diff changeset
   184
      | asts_of (Parser.Node ("_idtypdummy", pts as [Parser.Tip tok, _, _])) =
72e83c82c1d4 position constraint for bound dummy -- more PIDE markup;
wenzelm
parents: 52162
diff changeset
   185
          [Ast.Appl (Ast.Constant "_constrain" :: ast_of_dummy "_idtdummy" tok :: maps asts_of pts)]
42282
4fa41e068ff0 report literal tokens according to parsetree head;
wenzelm
parents: 42281
diff changeset
   186
      | asts_of (Parser.Node (a, pts)) =
4fa41e068ff0 report literal tokens according to parsetree head;
wenzelm
parents: 42281
diff changeset
   187
          let
4fa41e068ff0 report literal tokens according to parsetree head;
wenzelm
parents: 42281
diff changeset
   188
            val _ = pts |> List.app
4fa41e068ff0 report literal tokens according to parsetree head;
wenzelm
parents: 42281
diff changeset
   189
              (fn Parser.Node _ => () | Parser.Tip tok =>
4fa41e068ff0 report literal tokens according to parsetree head;
wenzelm
parents: 42281
diff changeset
   190
                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
   191
                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
   192
          in [trans a (maps asts_of pts)] end
49674
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 49662
diff changeset
   193
      | asts_of (Parser.Tip tok) = asts_of_token tok
42282
4fa41e068ff0 report literal tokens according to parsetree head;
wenzelm
parents: 42281
diff changeset
   194
4fa41e068ff0 report literal tokens according to parsetree head;
wenzelm
parents: 42281
diff changeset
   195
    and ast_of pt =
4fa41e068ff0 report literal tokens according to parsetree head;
wenzelm
parents: 42281
diff changeset
   196
      (case asts_of pt of
4fa41e068ff0 report literal tokens according to parsetree head;
wenzelm
parents: 42281
diff changeset
   197
        [ast] => ast
4fa41e068ff0 report literal tokens according to parsetree head;
wenzelm
parents: 42281
diff changeset
   198
      | 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
   199
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   200
    val ast = Exn.interruptible_capture ast_of parsetree;
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   201
  in (! reports, ast) end;
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   202
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
(* ast_to_term *)
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   205
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   206
fun ast_to_term ctxt trf =
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   207
  let
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   208
    fun trans a args =
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   209
      (case trf a of
42476
d0bc1268ef09 clarified auxiliary structure Lexicon.Syntax;
wenzelm
parents: 42470
diff changeset
   210
        NONE => Term.list_comb (Syntax.const a, args)
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   211
      | SOME f => f ctxt args);
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   212
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   213
    fun term_of (Ast.Constant a) = trans a []
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   214
      | term_of (Ast.Variable x) = Lexicon.read_var x
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   215
      | term_of (Ast.Appl (Ast.Constant a :: (asts as _ :: _))) =
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   216
          trans a (map term_of asts)
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   217
      | term_of (Ast.Appl (ast :: (asts as _ :: _))) =
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   218
          Term.list_comb (term_of ast, map term_of asts)
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   219
      | 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
   220
  in term_of end;
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   221
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   222
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   223
(* decode_term -- transform parse tree into raw term *)
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   224
44736
c2a3f1c84179 bulk reports for improved message throughput;
wenzelm
parents: 44735
diff changeset
   225
fun decode_term _ (result as (_: Position.report list, Exn.Exn _)) = result
43761
e72ba84ae58f tuned signature -- corresponding to Scala version;
wenzelm
parents: 43731
diff changeset
   226
  | decode_term ctxt (reports0, Exn.Res tm) =
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   227
      let
42250
cc5ac538f991 eliminated odd object-oriented type_context/term_context;
wenzelm
parents: 42249
diff changeset
   228
        fun get_const a =
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 42359
diff changeset
   229
          ((true, #1 (Term.dest_Const (Proof_Context.read_const_proper ctxt false a)))
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 42359
diff changeset
   230
            handle ERROR _ => (false, Consts.intern (Proof_Context.consts_of ctxt) a));
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 42359
diff changeset
   231
        val get_free = Proof_Context.intern_skolem ctxt;
42250
cc5ac538f991 eliminated odd object-oriented type_context/term_context;
wenzelm
parents: 42249
diff changeset
   232
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   233
        val reports = Unsynchronized.ref reports0;
44735
66862d02678c tuned signature;
wenzelm
parents: 44564
diff changeset
   234
        fun report ps = Position.store_reports reports ps;
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   235
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   236
        fun decode ps qs bs (Const ("_constrain", _) $ t $ typ) =
42264
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents: 42255
diff changeset
   237
              (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
   238
                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
   239
              | 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
   240
          | decode ps qs bs (Const ("_constrainAbs", _) $ t $ typ) =
42264
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents: 42255
diff changeset
   241
              (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
   242
                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
   243
              | 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
   244
          | decode _ qs bs (Abs (x, T, t)) =
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   245
              let
45412
7797f5351ec4 entity markup for bound variables;
wenzelm
parents: 44736
diff changeset
   246
                val id = serial ();
7797f5351ec4 entity markup for bound variables;
wenzelm
parents: 44736
diff changeset
   247
                val _ = report qs (markup_bound true qs) (x, id);
7797f5351ec4 entity markup for bound variables;
wenzelm
parents: 44736
diff changeset
   248
              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
   249
          | decode _ _ bs (t $ u) = decode [] [] bs t $ decode [] [] bs u
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   250
          | decode ps _ _ (Const (a, T)) =
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   251
              (case try Lexicon.unmark_fixed a of
42282
4fa41e068ff0 report literal tokens according to parsetree head;
wenzelm
parents: 42281
diff changeset
   252
                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
   253
              | NONE =>
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   254
                  let
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   255
                    val c =
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   256
                      (case try Lexicon.unmark_const a of
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   257
                        SOME c => c
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   258
                      | NONE => snd (get_const a));
42282
4fa41e068ff0 report literal tokens according to parsetree head;
wenzelm
parents: 42281
diff changeset
   259
                    val _ = report ps (markup_const ctxt) c;
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   260
                  in Const (c, T) end)
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   261
          | decode ps _ _ (Free (a, T)) =
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   262
              (case (get_free a, get_const a) of
42282
4fa41e068ff0 report literal tokens according to parsetree head;
wenzelm
parents: 42281
diff changeset
   263
                (SOME x, _) => (report ps (markup_free ctxt) x; Free (x, T))
4fa41e068ff0 report literal tokens according to parsetree head;
wenzelm
parents: 42281
diff changeset
   264
              | (_, (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
   265
              | (_, (false, c)) =>
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   266
                  if Long_Name.is_qualified c
42282
4fa41e068ff0 report literal tokens according to parsetree head;
wenzelm
parents: 42281
diff changeset
   267
                  then (report ps (markup_const ctxt) c; Const (c, T))
4fa41e068ff0 report literal tokens according to parsetree head;
wenzelm
parents: 42281
diff changeset
   268
                  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
   269
          | 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
   270
          | decode ps _ bs (t as Bound i) =
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   271
              (case try (nth bs) i of
45412
7797f5351ec4 entity markup for bound variables;
wenzelm
parents: 44736
diff changeset
   272
                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
   273
              | NONE => t);
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   274
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   275
        val tm' = Exn.interruptible_capture (fn () => decode [] [] [] tm) ();
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   276
      in (! reports, tm') end;
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   277
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   278
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
(** parse **)
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   281
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   282
(* results *)
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   283
43761
e72ba84ae58f tuned signature -- corresponding to Scala version;
wenzelm
parents: 43731
diff changeset
   284
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
   285
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
   286
46512
4f9f61f9b535 simplified configuration options for syntax ambiguity;
wenzelm
parents: 46506
diff changeset
   287
fun report_result ctxt pos ambig_msgs results =
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   288
  (case (proper_results results, failed_results results) of
44736
c2a3f1c84179 bulk reports for improved message throughput;
wenzelm
parents: 44735
diff changeset
   289
    ([], (reports, exn) :: _) => (Context_Position.reports ctxt reports; reraise exn)
c2a3f1c84179 bulk reports for improved message throughput;
wenzelm
parents: 44735
diff changeset
   290
  | ([(reports, x)], _) => (Context_Position.reports ctxt reports; x)
46512
4f9f61f9b535 simplified configuration options for syntax ambiguity;
wenzelm
parents: 46506
diff changeset
   291
  | _ =>
4f9f61f9b535 simplified configuration options for syntax ambiguity;
wenzelm
parents: 46506
diff changeset
   292
      if null ambig_msgs then
48992
0518bf89c777 renamed Position.str_of to Position.here;
wenzelm
parents: 48768
diff changeset
   293
        error ("Parse error: ambiguous syntax" ^ Position.here pos)
46512
4f9f61f9b535 simplified configuration options for syntax ambiguity;
wenzelm
parents: 46506
diff changeset
   294
      else error (cat_lines ambig_msgs));
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   295
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   296
42249
12a073670584 simplified standard parse/unparse;
wenzelm
parents: 42248
diff changeset
   297
(* parse raw asts *)
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   298
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   299
fun parse_asts ctxt raw root (syms, pos) =
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   300
  let
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 42359
diff changeset
   301
    val syn = Proof_Context.syn_of ctxt;
42253
c539d3c9c023 more abstract syntax translation;
wenzelm
parents: 42252
diff changeset
   302
    val ast_tr = Syntax.parse_ast_translation syn;
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   303
42251
050cc12dd985 explicit Syntax.tokenize, Syntax.parse;
wenzelm
parents: 42250
diff changeset
   304
    val toks = Syntax.tokenize syn raw syms;
44736
c2a3f1c84179 bulk reports for improved message throughput;
wenzelm
parents: 44735
diff changeset
   305
    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
   306
45641
20ef9135a9fb removed obsolete argument (cf. 954e9d6782ea);
wenzelm
parents: 45454
diff changeset
   307
    val pts = Syntax.parse syn root (filter Lexicon.is_proper toks)
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   308
      handle ERROR msg =>
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   309
        error (msg ^
50201
c26369c9eda6 Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents: 49690
diff changeset
   310
          implode (map (Markup.markup Markup.report o Lexicon.reported_token_range ctxt) toks));
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   311
    val len = length pts;
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   312
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   313
    val limit = Config.get ctxt Syntax.ambiguity_limit;
46512
4f9f61f9b535 simplified configuration options for syntax ambiguity;
wenzelm
parents: 46506
diff changeset
   314
    val ambig_msgs =
4f9f61f9b535 simplified configuration options for syntax ambiguity;
wenzelm
parents: 46506
diff changeset
   315
      if len <= 1 then []
46506
c7faa011bfa7 simplified configuration options for syntax ambiguity;
wenzelm
parents: 46126
diff changeset
   316
      else
46512
4f9f61f9b535 simplified configuration options for syntax ambiguity;
wenzelm
parents: 46506
diff changeset
   317
        [cat_lines
48992
0518bf89c777 renamed Position.str_of to Position.here;
wenzelm
parents: 48768
diff changeset
   318
          (("Ambiguous input" ^ Position.here (Position.reset_range pos) ^
46512
4f9f61f9b535 simplified configuration options for syntax ambiguity;
wenzelm
parents: 46506
diff changeset
   319
            "\nproduces " ^ string_of_int len ^ " parse trees" ^
4f9f61f9b535 simplified configuration options for syntax ambiguity;
wenzelm
parents: 46506
diff changeset
   320
            (if len <= limit then "" else " (" ^ string_of_int limit ^ " displayed)") ^ ":") ::
52517
89c5af70553f tuned message;
wenzelm
parents: 52464
diff changeset
   321
            map (Pretty.string_of o Pretty.item o single o Parser.pretty_parsetree)
89c5af70553f tuned message;
wenzelm
parents: 52464
diff changeset
   322
              (take limit pts))];
46512
4f9f61f9b535 simplified configuration options for syntax ambiguity;
wenzelm
parents: 46506
diff changeset
   323
52188
2da0033370a0 report markup for ast translations;
wenzelm
parents: 52186
diff changeset
   324
  in (ambig_msgs, map (parsetree_to_ast ctxt ast_tr) pts) end;
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   325
46506
c7faa011bfa7 simplified configuration options for syntax ambiguity;
wenzelm
parents: 46126
diff changeset
   326
fun parse_tree ctxt root input =
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   327
  let
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 42359
diff changeset
   328
    val syn = Proof_Context.syn_of ctxt;
42253
c539d3c9c023 more abstract syntax translation;
wenzelm
parents: 42252
diff changeset
   329
    val tr = Syntax.parse_translation syn;
42255
097c93dcd877 eliminated slightly odd Syntax.rep_syntax;
wenzelm
parents: 42254
diff changeset
   330
    val parse_rules = Syntax.parse_rules syn;
46512
4f9f61f9b535 simplified configuration options for syntax ambiguity;
wenzelm
parents: 46506
diff changeset
   331
    val (ambig_msgs, asts) = parse_asts ctxt false root input;
4f9f61f9b535 simplified configuration options for syntax ambiguity;
wenzelm
parents: 46506
diff changeset
   332
    val results =
4f9f61f9b535 simplified configuration options for syntax ambiguity;
wenzelm
parents: 46506
diff changeset
   333
      (map o apsnd o Exn.maps_result)
4f9f61f9b535 simplified configuration options for syntax ambiguity;
wenzelm
parents: 46506
diff changeset
   334
        (Ast.normalize ctxt parse_rules #> Exn.interruptible_capture (ast_to_term ctxt tr)) asts;
4f9f61f9b535 simplified configuration options for syntax ambiguity;
wenzelm
parents: 46506
diff changeset
   335
  in (ambig_msgs, results) end;
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   336
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   337
42249
12a073670584 simplified standard parse/unparse;
wenzelm
parents: 42248
diff changeset
   338
(* parse logical entities *)
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   339
42241
dd8029f71e1c separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff changeset
   340
fun parse_failed ctxt pos msg kind =
dd8029f71e1c separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff changeset
   341
  cat_error msg ("Failed to parse " ^ kind ^
50201
c26369c9eda6 Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents: 49690
diff changeset
   342
    Markup.markup Markup.report (Context_Position.reported_text ctxt pos Markup.bad ""));
42241
dd8029f71e1c separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff changeset
   343
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
fun parse_sort ctxt =
50201
c26369c9eda6 Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents: 49690
diff changeset
   345
  Syntax.parse_token ctxt Term_XML.Decode.sort Markup.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
   346
    (fn (syms, pos) =>
46506
c7faa011bfa7 simplified configuration options for syntax ambiguity;
wenzelm
parents: 46126
diff changeset
   347
      parse_tree ctxt "sort" (syms, pos)
46512
4f9f61f9b535 simplified configuration options for syntax ambiguity;
wenzelm
parents: 46506
diff changeset
   348
      |> uncurry (report_result ctxt pos)
45427
fca432074fb2 sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents: 45423
diff changeset
   349
      |> 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
   350
      |> 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
   351
      handle ERROR msg => parse_failed ctxt pos msg "sort");
42241
dd8029f71e1c separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff changeset
   352
43731
70072780e095 inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents: 43552
diff changeset
   353
fun parse_typ ctxt =
50201
c26369c9eda6 Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents: 49690
diff changeset
   354
  Syntax.parse_token ctxt Term_XML.Decode.typ Markup.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
   355
    (fn (syms, pos) =>
46506
c7faa011bfa7 simplified configuration options for syntax ambiguity;
wenzelm
parents: 46126
diff changeset
   356
      parse_tree ctxt "type" (syms, pos)
46512
4f9f61f9b535 simplified configuration options for syntax ambiguity;
wenzelm
parents: 46506
diff changeset
   357
      |> uncurry (report_result ctxt pos)
45427
fca432074fb2 sort assignment before simultaneous term_check, not isolated parse_term;
wenzelm
parents: 45423
diff changeset
   358
      |> 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
   359
      handle ERROR msg => parse_failed ctxt pos msg "type");
42241
dd8029f71e1c separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff changeset
   360
43731
70072780e095 inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents: 43552
diff changeset
   361
fun parse_term is_prop ctxt =
42241
dd8029f71e1c separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff changeset
   362
  let
42281
69d4543811d0 simplified read_term vs. read_prop;
wenzelm
parents: 42267
diff changeset
   363
    val (markup, kind, root, constrain) =
69d4543811d0 simplified read_term vs. read_prop;
wenzelm
parents: 42267
diff changeset
   364
      if is_prop
50201
c26369c9eda6 Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents: 49690
diff changeset
   365
      then (Markup.prop, "proposition", "prop", Type.constraint propT)
c26369c9eda6 Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents: 49690
diff changeset
   366
      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
   367
    val decode = constrain o Term_XML.Decode.term;
42249
12a073670584 simplified standard parse/unparse;
wenzelm
parents: 42248
diff changeset
   368
  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
   369
    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
   370
      (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
   371
        let
46512
4f9f61f9b535 simplified configuration options for syntax ambiguity;
wenzelm
parents: 46506
diff changeset
   372
          val (ambig_msgs, results) = parse_tree ctxt root (syms, pos) ||> map (decode_term ctxt);
46506
c7faa011bfa7 simplified configuration options for syntax ambiguity;
wenzelm
parents: 46126
diff changeset
   373
          val parsed_len = length (proper_results results);
42249
12a073670584 simplified standard parse/unparse;
wenzelm
parents: 42248
diff changeset
   374
46512
4f9f61f9b535 simplified configuration options for syntax ambiguity;
wenzelm
parents: 46506
diff changeset
   375
          val ambiguity_warning = Config.get ctxt Syntax.ambiguity_warning;
43731
70072780e095 inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents: 43552
diff changeset
   376
          val limit = Config.get ctxt Syntax.ambiguity_limit;
42249
12a073670584 simplified standard parse/unparse;
wenzelm
parents: 42248
diff changeset
   377
43731
70072780e095 inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents: 43552
diff changeset
   378
          (*brute-force disambiguation via type-inference*)
43761
e72ba84ae58f tuned signature -- corresponding to Scala version;
wenzelm
parents: 43731
diff changeset
   379
          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
   380
            handle exn as ERROR _ => Exn.Exn exn;
42249
12a073670584 simplified standard parse/unparse;
wenzelm
parents: 42248
diff changeset
   381
43731
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 results' =
46506
c7faa011bfa7 simplified configuration options for syntax ambiguity;
wenzelm
parents: 46126
diff changeset
   383
            if parsed_len > 1 then
46989
88b0a8052c75 added Syntax.read_typs;
wenzelm
parents: 46512
diff changeset
   384
              (grouped 10 (Par_List.map_name "Syntax_Phases.parse_term") o apsnd o Exn.maps_result)
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
                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
   386
            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
   387
          val reports' = fst (hd results');
42249
12a073670584 simplified standard parse/unparse;
wenzelm
parents: 42248
diff changeset
   388
43731
70072780e095 inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents: 43552
diff changeset
   389
          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
   390
          val checked = map snd (proper_results results');
46506
c7faa011bfa7 simplified configuration options for syntax ambiguity;
wenzelm
parents: 46126
diff changeset
   391
          val checked_len = length checked;
42249
12a073670584 simplified standard parse/unparse;
wenzelm
parents: 42248
diff changeset
   392
52517
89c5af70553f tuned message;
wenzelm
parents: 52464
diff changeset
   393
          val pretty_term = Syntax.pretty_term (Config.put Printer.show_brackets true ctxt);
43731
70072780e095 inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents: 43552
diff changeset
   394
        in
46506
c7faa011bfa7 simplified configuration options for syntax ambiguity;
wenzelm
parents: 46126
diff changeset
   395
          if checked_len = 0 then
46512
4f9f61f9b535 simplified configuration options for syntax ambiguity;
wenzelm
parents: 46506
diff changeset
   396
            report_result ctxt pos []
4f9f61f9b535 simplified configuration options for syntax ambiguity;
wenzelm
parents: 46506
diff changeset
   397
              [(reports', Exn.Exn (Exn.EXCEPTIONS (map ERROR ambig_msgs @ errs)))]
46506
c7faa011bfa7 simplified configuration options for syntax ambiguity;
wenzelm
parents: 46126
diff changeset
   398
          else if checked_len = 1 then
46512
4f9f61f9b535 simplified configuration options for syntax ambiguity;
wenzelm
parents: 46506
diff changeset
   399
            (if parsed_len > 1 andalso ambiguity_warning 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
   400
              Context_Position.if_visible ctxt warning
46512
4f9f61f9b535 simplified configuration options for syntax ambiguity;
wenzelm
parents: 46506
diff changeset
   401
                (cat_lines (ambig_msgs @
4f9f61f9b535 simplified configuration options for syntax ambiguity;
wenzelm
parents: 46506
diff changeset
   402
                  ["Fortunately, only one parse tree is type correct" ^
48992
0518bf89c777 renamed Position.str_of to Position.here;
wenzelm
parents: 48768
diff changeset
   403
                  Position.here (Position.reset_range pos) ^
46512
4f9f61f9b535 simplified configuration options for syntax ambiguity;
wenzelm
parents: 46506
diff changeset
   404
                  ",\nbut you may still want to disambiguate your grammar or your input."]))
4f9f61f9b535 simplified configuration options for syntax ambiguity;
wenzelm
parents: 46506
diff changeset
   405
             else (); report_result ctxt pos [] results')
43731
70072780e095 inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents: 43552
diff changeset
   406
          else
46512
4f9f61f9b535 simplified configuration options for syntax ambiguity;
wenzelm
parents: 46506
diff changeset
   407
            report_result ctxt pos []
4f9f61f9b535 simplified configuration options for syntax ambiguity;
wenzelm
parents: 46506
diff changeset
   408
              [(reports', Exn.Exn (ERROR (cat_lines (ambig_msgs @
4f9f61f9b535 simplified configuration options for syntax ambiguity;
wenzelm
parents: 46506
diff changeset
   409
                (("Ambiguous input\n" ^ string_of_int checked_len ^ " terms are type correct" ^
46506
c7faa011bfa7 simplified configuration options for syntax ambiguity;
wenzelm
parents: 46126
diff changeset
   410
                  (if checked_len <= limit then ""
c7faa011bfa7 simplified configuration options for syntax ambiguity;
wenzelm
parents: 46126
diff changeset
   411
                   else " (" ^ string_of_int limit ^ " displayed)") ^ ":") ::
52517
89c5af70553f tuned message;
wenzelm
parents: 52464
diff changeset
   412
                  map (Pretty.string_of o Pretty.item o single o pretty_term)
89c5af70553f tuned message;
wenzelm
parents: 52464
diff changeset
   413
                    (take limit checked))))))]
43731
70072780e095 inner syntax supports inlined YXML according to Term_XML (particularly useful for producing text under program control);
wenzelm
parents: 43552
diff changeset
   414
        end handle ERROR msg => parse_failed ctxt pos msg kind)
42249
12a073670584 simplified standard parse/unparse;
wenzelm
parents: 42248
diff changeset
   415
  end;
42241
dd8029f71e1c separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff changeset
   416
dd8029f71e1c separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff changeset
   417
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   418
(* parse_ast_pattern *)
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
fun parse_ast_pattern ctxt (root, str) =
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   421
  let
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 42359
diff changeset
   422
    val syn = Proof_Context.syn_of ctxt;
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   423
52188
2da0033370a0 report markup for ast translations;
wenzelm
parents: 52186
diff changeset
   424
    val reports = Unsynchronized.ref ([]: Position.report list);
2da0033370a0 report markup for ast translations;
wenzelm
parents: 52186
diff changeset
   425
    fun report ps = Position.store_reports reports ps;
2da0033370a0 report markup for ast translations;
wenzelm
parents: 52186
diff changeset
   426
2da0033370a0 report markup for ast translations;
wenzelm
parents: 52186
diff changeset
   427
    fun decode_const ps c = (report ps (markup_entity ctxt) c; Ast.Constant c);
2da0033370a0 report markup for ast translations;
wenzelm
parents: 52186
diff changeset
   428
    fun decode_var ps x = (report ps (fn () => [Markup.name x Markup.free]) (); Ast.Variable x);
2da0033370a0 report markup for ast translations;
wenzelm
parents: 52186
diff changeset
   429
    fun decode_appl ps asts = Ast.Appl (map (decode ps) asts)
2da0033370a0 report markup for ast translations;
wenzelm
parents: 52186
diff changeset
   430
    and decode ps (Ast.Constant c) = decode_const ps c
2da0033370a0 report markup for ast translations;
wenzelm
parents: 52186
diff changeset
   431
      | decode ps (Ast.Variable x) =
42298
d622145603ee more accurate markup for syntax consts, notably binders which point back to the original logical entity;
wenzelm
parents: 42296
diff changeset
   432
          if is_some (Syntax.lookup_const syn x) orelse Long_Name.is_qualified x
52188
2da0033370a0 report markup for ast translations;
wenzelm
parents: 52186
diff changeset
   433
          then decode_const ps x
2da0033370a0 report markup for ast translations;
wenzelm
parents: 52186
diff changeset
   434
          else decode_var ps x
2da0033370a0 report markup for ast translations;
wenzelm
parents: 52186
diff changeset
   435
      | decode ps (Ast.Appl (asts as (Ast.Constant c :: ast :: Ast.Variable x :: args))) =
2da0033370a0 report markup for ast translations;
wenzelm
parents: 52186
diff changeset
   436
          if member (op =) Term_Position.markers c then
2da0033370a0 report markup for ast translations;
wenzelm
parents: 52186
diff changeset
   437
            (case Term_Position.decode x of
2da0033370a0 report markup for ast translations;
wenzelm
parents: 52186
diff changeset
   438
              SOME p => Ast.mk_appl (decode (p :: ps) ast) (map (decode ps) args)
2da0033370a0 report markup for ast translations;
wenzelm
parents: 52186
diff changeset
   439
            | NONE => decode_appl ps asts)
2da0033370a0 report markup for ast translations;
wenzelm
parents: 52186
diff changeset
   440
          else decode_appl ps asts
2da0033370a0 report markup for ast translations;
wenzelm
parents: 52186
diff changeset
   441
      | decode ps (Ast.Appl asts) = decode_appl ps asts;
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   442
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   443
    val (syms, pos) = Syntax.read_token str;
52188
2da0033370a0 report markup for ast translations;
wenzelm
parents: 52186
diff changeset
   444
    val ast =
2da0033370a0 report markup for ast translations;
wenzelm
parents: 52186
diff changeset
   445
      parse_asts ctxt true root (syms, pos)
2da0033370a0 report markup for ast translations;
wenzelm
parents: 52186
diff changeset
   446
      |> uncurry (report_result ctxt pos)
2da0033370a0 report markup for ast translations;
wenzelm
parents: 52186
diff changeset
   447
      |> decode [];
2da0033370a0 report markup for ast translations;
wenzelm
parents: 52186
diff changeset
   448
    val _ = Context_Position.reports ctxt (! reports);
2da0033370a0 report markup for ast translations;
wenzelm
parents: 52186
diff changeset
   449
  in ast end;
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   450
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   451
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   452
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   453
(** encode parse trees **)
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   454
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   455
(* term_of_sort *)
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_sort S =
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   458
  let
42476
d0bc1268ef09 clarified auxiliary structure Lexicon.Syntax;
wenzelm
parents: 42470
diff changeset
   459
    val class = Syntax.const o Lexicon.mark_class;
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   460
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   461
    fun classes [c] = class c
42476
d0bc1268ef09 clarified auxiliary structure Lexicon.Syntax;
wenzelm
parents: 42470
diff changeset
   462
      | classes (c :: cs) = Syntax.const "_classes" $ class c $ classes cs;
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   463
  in
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   464
    (case S of
42476
d0bc1268ef09 clarified auxiliary structure Lexicon.Syntax;
wenzelm
parents: 42470
diff changeset
   465
      [] => Syntax.const "_topsort"
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   466
    | [c] => class c
42476
d0bc1268ef09 clarified auxiliary structure Lexicon.Syntax;
wenzelm
parents: 42470
diff changeset
   467
    | cs => Syntax.const "_sort" $ classes cs)
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   468
  end;
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   469
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   470
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   471
(* term_of_typ *)
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   472
42247
12fe41a92cd5 typed_print_translation: discontinued show_sorts argument;
wenzelm
parents: 42245
diff changeset
   473
fun term_of_typ ctxt ty =
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   474
  let
52210
0226035df99d more explicit Printer.type_emphasis, depending on show_type_emphasis;
wenzelm
parents: 52188
diff changeset
   475
    val show_sorts = Config.get ctxt show_sorts orelse Config.get ctxt show_markup;
42247
12fe41a92cd5 typed_print_translation: discontinued show_sorts argument;
wenzelm
parents: 42245
diff changeset
   476
49686
678aa92e921c printed "_ofsort" is subject to show_markup as well;
wenzelm
parents: 49685
diff changeset
   477
    fun ofsort t raw_S =
52210
0226035df99d more explicit Printer.type_emphasis, depending on show_type_emphasis;
wenzelm
parents: 52188
diff changeset
   478
      if show_sorts then
49686
678aa92e921c printed "_ofsort" is subject to show_markup as well;
wenzelm
parents: 49685
diff changeset
   479
        let val S = #2 (Term_Position.decode_positionS raw_S)
678aa92e921c printed "_ofsort" is subject to show_markup as well;
wenzelm
parents: 49685
diff changeset
   480
        in if S = dummyS then t else Syntax.const "_ofsort" $ t $ term_of_sort S end
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   481
      else t;
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   482
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   483
    fun term_of (Type (a, Ts)) =
42476
d0bc1268ef09 clarified auxiliary structure Lexicon.Syntax;
wenzelm
parents: 42470
diff changeset
   484
          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
   485
      | term_of (TFree (x, S)) =
42290
b1f544c84040 discontinued special treatment of structure Lexicon;
wenzelm
parents: 42289
diff changeset
   486
          if is_some (Term_Position.decode x) then Syntax.free x
49686
678aa92e921c printed "_ofsort" is subject to show_markup as well;
wenzelm
parents: 49685
diff changeset
   487
          else ofsort (Syntax.const "_tfree" $ Syntax.free x) S
678aa92e921c printed "_ofsort" is subject to show_markup as well;
wenzelm
parents: 49685
diff changeset
   488
      | term_of (TVar (xi, S)) = ofsort (Syntax.const "_tvar" $ Syntax.var xi) S;
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   489
  in term_of ty end;
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   490
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   491
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   492
(* simple_ast_of *)
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   493
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   494
fun simple_ast_of ctxt =
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   495
  let
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   496
    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
   497
    fun ast_of (Const (c, _)) = Ast.Constant c
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   498
      | ast_of (Free (x, _)) = Ast.Variable x
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   499
      | 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
   500
      | ast_of (t as _ $ _) =
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   501
          let val (f, args) = strip_comb t
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   502
          in Ast.mk_appl (ast_of f) (map ast_of args) end
42408
282b7a3065d3 explicit markup for loose bounds;
wenzelm
parents: 42382
diff changeset
   503
      | 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
   504
      | ast_of (Abs _) = raise Fail "simple_ast_of: Abs";
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   505
  in ast_of end;
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   506
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   507
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   508
(* sort_to_ast and typ_to_ast *)
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   509
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   510
fun ast_of_termT ctxt trf tm =
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   511
  let
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   512
    val ctxt' = Config.put show_sorts false ctxt;
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   513
    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
   514
      | 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
   515
      | ast_of (Const (a, _)) = trans a []
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   516
      | ast_of (t as _ $ _) =
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   517
          (case strip_comb t of
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   518
            (Const (a, _), args) => trans a args
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   519
          | (f, args) => Ast.Appl (map ast_of (f :: args)))
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   520
      | ast_of t = simple_ast_of ctxt t
42254
f427c9890c46 more abstract print translation;
wenzelm
parents: 42253
diff changeset
   521
    and trans a args = ast_of (trf a ctxt' dummyT args)
f427c9890c46 more abstract print translation;
wenzelm
parents: 42253
diff changeset
   522
      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
   523
  in ast_of tm end;
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   524
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   525
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
   526
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
   527
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   528
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   529
(* term_to_ast *)
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   530
52177
wenzelm
parents: 52163
diff changeset
   531
local
wenzelm
parents: 52163
diff changeset
   532
wenzelm
parents: 52163
diff changeset
   533
fun mark_aprop tm =
wenzelm
parents: 52163
diff changeset
   534
  let
wenzelm
parents: 52163
diff changeset
   535
    fun aprop t = Syntax.const "_aprop" $ t;
wenzelm
parents: 52163
diff changeset
   536
wenzelm
parents: 52163
diff changeset
   537
    fun is_prop Ts t =
52464
36497ee02ed8 clarified printing of type constraints vs. annotations, notably in goal display with its Type_Annotation.ignore_free_types;
wenzelm
parents: 52219
diff changeset
   538
      Type_Annotation.clean (Type_Annotation.fastype_of Ts t) = propT
36497ee02ed8 clarified printing of type constraints vs. annotations, notably in goal display with its Type_Annotation.ignore_free_types;
wenzelm
parents: 52219
diff changeset
   539
        handle TERM _ => false;
52177
wenzelm
parents: 52163
diff changeset
   540
wenzelm
parents: 52163
diff changeset
   541
    fun is_term (Const ("Pure.term", _) $ _) = true
wenzelm
parents: 52163
diff changeset
   542
      | is_term _ = false;
wenzelm
parents: 52163
diff changeset
   543
wenzelm
parents: 52163
diff changeset
   544
    fun mark _ (t as Const _) = t
wenzelm
parents: 52163
diff changeset
   545
      | mark Ts (t as Const ("_bound", _) $ u) = if is_prop Ts u then aprop t else t
52464
36497ee02ed8 clarified printing of type constraints vs. annotations, notably in goal display with its Type_Annotation.ignore_free_types;
wenzelm
parents: 52219
diff changeset
   546
      | mark Ts (t as Free _) = if is_prop Ts t then aprop t else t
36497ee02ed8 clarified printing of type constraints vs. annotations, notably in goal display with its Type_Annotation.ignore_free_types;
wenzelm
parents: 52219
diff changeset
   547
      | mark Ts (t as Var _) = if is_prop Ts t then aprop t else t
52177
wenzelm
parents: 52163
diff changeset
   548
      | mark Ts (t as Bound _) = if is_prop Ts t then aprop t else t
wenzelm
parents: 52163
diff changeset
   549
      | mark Ts (Abs (x, T, t)) = Abs (x, T, mark (T :: Ts) t)
wenzelm
parents: 52163
diff changeset
   550
      | mark Ts (t as t1 $ (t2 as Const ("TYPE", Type ("itself", [T])))) =
wenzelm
parents: 52163
diff changeset
   551
          if is_prop Ts t andalso not (is_term t) then Const ("_type_prop", T) $ mark Ts t1
wenzelm
parents: 52163
diff changeset
   552
          else mark Ts t1 $ mark Ts t2
wenzelm
parents: 52163
diff changeset
   553
      | mark Ts (t as t1 $ t2) =
wenzelm
parents: 52163
diff changeset
   554
          (if is_Const (Term.head_of t) orelse not (is_prop Ts t) then I else aprop)
wenzelm
parents: 52163
diff changeset
   555
            (mark Ts t1 $ mark Ts t2);
wenzelm
parents: 52163
diff changeset
   556
  in mark [] tm end;
wenzelm
parents: 52163
diff changeset
   557
52186
wenzelm
parents: 52185
diff changeset
   558
fun prune_types ctxt tm =
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   559
  let
52211
66bc827e37f8 explicit support for type annotations within printed syntax trees;
wenzelm
parents: 52210
diff changeset
   560
    fun regard t t' seen =
66bc827e37f8 explicit support for type annotations within printed syntax trees;
wenzelm
parents: 52210
diff changeset
   561
      if Type_Annotation.is_omitted (Type_Annotation.fastype_of [] t) then (t, seen)
66bc827e37f8 explicit support for type annotations within printed syntax trees;
wenzelm
parents: 52210
diff changeset
   562
      else if member (op aconv) seen t then (t', seen)
66bc827e37f8 explicit support for type annotations within printed syntax trees;
wenzelm
parents: 52210
diff changeset
   563
      else (t, t :: seen);
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   564
52211
66bc827e37f8 explicit support for type annotations within printed syntax trees;
wenzelm
parents: 52210
diff changeset
   565
    fun prune (t as Const _, seen) = (t, seen)
66bc827e37f8 explicit support for type annotations within printed syntax trees;
wenzelm
parents: 52210
diff changeset
   566
      | prune (t as Free (x, T), seen) = regard t (Free (x, Type_Annotation.ignore_type T)) seen
66bc827e37f8 explicit support for type annotations within printed syntax trees;
wenzelm
parents: 52210
diff changeset
   567
      | prune (t as Var (xi, T), seen) = regard t (Var (xi, Type_Annotation.ignore_type T)) seen
66bc827e37f8 explicit support for type annotations within printed syntax trees;
wenzelm
parents: 52210
diff changeset
   568
      | prune (t as Bound _, seen) = (t, seen)
66bc827e37f8 explicit support for type annotations within printed syntax trees;
wenzelm
parents: 52210
diff changeset
   569
      | prune (Abs (x, T, t), seen) =
52186
wenzelm
parents: 52185
diff changeset
   570
          let val (t', seen') = prune (t, seen);
52211
66bc827e37f8 explicit support for type annotations within printed syntax trees;
wenzelm
parents: 52210
diff changeset
   571
          in (Abs (x, T, t'), seen') end
52186
wenzelm
parents: 52185
diff changeset
   572
      | prune (t1 $ t2, seen) =
52177
wenzelm
parents: 52163
diff changeset
   573
          let
52186
wenzelm
parents: 52185
diff changeset
   574
            val (t1', seen') = prune (t1, seen);
wenzelm
parents: 52185
diff changeset
   575
            val (t2', seen'') = prune (t2, seen');
52177
wenzelm
parents: 52163
diff changeset
   576
          in (t1' $ t2', seen'') end;
52186
wenzelm
parents: 52185
diff changeset
   577
  in #1 (prune (tm, [])) end;
52177
wenzelm
parents: 52163
diff changeset
   578
52186
wenzelm
parents: 52185
diff changeset
   579
fun mark_atoms {structs, fixes} is_syntax_const ctxt tm =
wenzelm
parents: 52185
diff changeset
   580
  let
wenzelm
parents: 52185
diff changeset
   581
    val show_structs = Config.get ctxt show_structs;
wenzelm
parents: 52185
diff changeset
   582
wenzelm
parents: 52185
diff changeset
   583
    fun mark ((t as Const (c, _)) $ u) =
52160
7746c9f1baf3 tuned signature;
wenzelm
parents: 52143
diff changeset
   584
          if member (op =) Pure_Thy.token_markers c
52186
wenzelm
parents: 52185
diff changeset
   585
          then t $ u else mark t $ mark u
wenzelm
parents: 52185
diff changeset
   586
      | mark (t $ u) = mark t $ mark u
wenzelm
parents: 52185
diff changeset
   587
      | mark (Abs (x, T, t)) = Abs (x, T, mark t)
wenzelm
parents: 52185
diff changeset
   588
      | mark (t as Const (c, T)) =
42252
wenzelm
parents: 42251
diff changeset
   589
          if is_syntax_const c then t
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   590
          else Const (Lexicon.mark_const c, T)
52186
wenzelm
parents: 52185
diff changeset
   591
      | mark (t as Free (x, T)) =
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   592
          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
   593
            if i = 0 andalso member (op =) fixes x then
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   594
              Const (Lexicon.mark_fixed x, T)
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   595
            else if i = 1 andalso not show_structs then
42476
d0bc1268ef09 clarified auxiliary structure Lexicon.Syntax;
wenzelm
parents: 42470
diff changeset
   596
              Syntax.const "_struct" $ Syntax.const "_indexdefault"
d0bc1268ef09 clarified auxiliary structure Lexicon.Syntax;
wenzelm
parents: 42470
diff changeset
   597
            else Syntax.const "_free" $ t
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   598
          end
52186
wenzelm
parents: 52185
diff changeset
   599
      | mark (t as Var (xi, T)) =
42288
2074b31650e6 discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
wenzelm
parents: 42284
diff changeset
   600
          if xi = Syntax_Ext.dddot_indexname then Const ("_DDDOT", T)
42476
d0bc1268ef09 clarified auxiliary structure Lexicon.Syntax;
wenzelm
parents: 42470
diff changeset
   601
          else Syntax.const "_var" $ t
52186
wenzelm
parents: 52185
diff changeset
   602
      | mark a = a;
wenzelm
parents: 52185
diff changeset
   603
  in mark tm end;
wenzelm
parents: 52185
diff changeset
   604
wenzelm
parents: 52185
diff changeset
   605
in
wenzelm
parents: 52185
diff changeset
   606
wenzelm
parents: 52185
diff changeset
   607
fun term_to_ast idents is_syntax_const ctxt trf tm =
wenzelm
parents: 52185
diff changeset
   608
  let
wenzelm
parents: 52185
diff changeset
   609
    val show_types = Config.get ctxt show_types orelse Config.get ctxt show_sorts;
wenzelm
parents: 52185
diff changeset
   610
    val show_markup = Config.get ctxt show_markup;
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   611
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   612
    fun ast_of tm =
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   613
      (case strip_comb tm of
42284
326f57825e1a explicit structure Syntax_Trans;
wenzelm
parents: 42282
diff changeset
   614
        (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
   615
      | ((c as Const ("_free", _)), Free (x, T) :: ts) =>
42290
b1f544c84040 discontinued special treatment of structure Lexicon;
wenzelm
parents: 42289
diff changeset
   616
          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
   617
      | ((c as Const ("_var", _)), Var (xi, T) :: ts) =>
42290
b1f544c84040 discontinued special treatment of structure Lexicon;
wenzelm
parents: 42289
diff changeset
   618
          Ast.mk_appl (constrain (c $ Syntax.var xi) T) (map ast_of ts)
49660
de49d9b4d7bc more explicit Syntax_Trans.mark_bound_abs/mark_bound_body: preserve type information for show_markup;
wenzelm
parents: 49659
diff changeset
   619
      | ((c as Const ("_bound", B)), Free (x, T) :: ts) =>
de49d9b4d7bc more explicit Syntax_Trans.mark_bound_abs/mark_bound_body: preserve type information for show_markup;
wenzelm
parents: 49659
diff changeset
   620
          let
de49d9b4d7bc more explicit Syntax_Trans.mark_bound_abs/mark_bound_body: preserve type information for show_markup;
wenzelm
parents: 49659
diff changeset
   621
            val X =
de49d9b4d7bc more explicit Syntax_Trans.mark_bound_abs/mark_bound_body: preserve type information for show_markup;
wenzelm
parents: 49659
diff changeset
   622
              if show_markup andalso not show_types orelse B <> dummyT then T
de49d9b4d7bc more explicit Syntax_Trans.mark_bound_abs/mark_bound_body: preserve type information for show_markup;
wenzelm
parents: 49659
diff changeset
   623
              else dummyT;
de49d9b4d7bc more explicit Syntax_Trans.mark_bound_abs/mark_bound_body: preserve type information for show_markup;
wenzelm
parents: 49659
diff changeset
   624
          in Ast.mk_appl (constrain (c $ Syntax.free x) X) (map ast_of ts) end
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   625
      | (Const ("_idtdummy", T), ts) =>
42476
d0bc1268ef09 clarified auxiliary structure Lexicon.Syntax;
wenzelm
parents: 42470
diff changeset
   626
          Ast.mk_appl (constrain (Syntax.const "_idtdummy") T) (map ast_of ts)
52219
c8ee9c0a3a64 observe type annotations in print translations as well, notably type_constraint_tr';
wenzelm
parents: 52211
diff changeset
   627
      | (const as Const (c, T), ts) => trans c (Type_Annotation.smash T) ts
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   628
      | (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
   629
42254
f427c9890c46 more abstract print translation;
wenzelm
parents: 42253
diff changeset
   630
    and trans a T args = ast_of (trf a ctxt T args)
f427c9890c46 more abstract print translation;
wenzelm
parents: 42253
diff changeset
   631
      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
   632
52211
66bc827e37f8 explicit support for type annotations within printed syntax trees;
wenzelm
parents: 52210
diff changeset
   633
    and constrain t T0 =
52464
36497ee02ed8 clarified printing of type constraints vs. annotations, notably in goal display with its Type_Annotation.ignore_free_types;
wenzelm
parents: 52219
diff changeset
   634
      let
36497ee02ed8 clarified printing of type constraints vs. annotations, notably in goal display with its Type_Annotation.ignore_free_types;
wenzelm
parents: 52219
diff changeset
   635
        val T =
36497ee02ed8 clarified printing of type constraints vs. annotations, notably in goal display with its Type_Annotation.ignore_free_types;
wenzelm
parents: 52219
diff changeset
   636
          if show_markup andalso not show_types
36497ee02ed8 clarified printing of type constraints vs. annotations, notably in goal display with its Type_Annotation.ignore_free_types;
wenzelm
parents: 52219
diff changeset
   637
          then Type_Annotation.clean T0
36497ee02ed8 clarified printing of type constraints vs. annotations, notably in goal display with its Type_Annotation.ignore_free_types;
wenzelm
parents: 52219
diff changeset
   638
          else Type_Annotation.smash T0;
36497ee02ed8 clarified printing of type constraints vs. annotations, notably in goal display with its Type_Annotation.ignore_free_types;
wenzelm
parents: 52219
diff changeset
   639
      in
52211
66bc827e37f8 explicit support for type annotations within printed syntax trees;
wenzelm
parents: 52210
diff changeset
   640
        if (show_types orelse show_markup) andalso T <> dummyT then
66bc827e37f8 explicit support for type annotations within printed syntax trees;
wenzelm
parents: 52210
diff changeset
   641
          Ast.Appl [Ast.Constant "_constrain", simple_ast_of ctxt t,
66bc827e37f8 explicit support for type annotations within printed syntax trees;
wenzelm
parents: 52210
diff changeset
   642
            ast_of_termT ctxt trf (term_of_typ ctxt T)]
66bc827e37f8 explicit support for type annotations within printed syntax trees;
wenzelm
parents: 52210
diff changeset
   643
        else simple_ast_of ctxt t
66bc827e37f8 explicit support for type annotations within printed syntax trees;
wenzelm
parents: 52210
diff changeset
   644
      end;
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   645
  in
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   646
    tm
52177
wenzelm
parents: 52163
diff changeset
   647
    |> mark_aprop
52186
wenzelm
parents: 52185
diff changeset
   648
    |> show_types ? prune_types ctxt
55014
a93f496f6c30 general notion of auxiliary bounds within context;
wenzelm
parents: 53171
diff changeset
   649
    |> Variable.revert_bounds ctxt
52186
wenzelm
parents: 52185
diff changeset
   650
    |> mark_atoms idents is_syntax_const ctxt
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   651
    |> ast_of
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   652
  end;
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   653
52177
wenzelm
parents: 52163
diff changeset
   654
end;
wenzelm
parents: 52163
diff changeset
   655
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   656
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   657
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   658
(** unparse **)
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   659
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   660
local
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   661
42267
9566078ad905 simplified printer context: uniform externing and static token translations;
wenzelm
parents: 42264
diff changeset
   662
fun free_or_skolem ctxt x =
9566078ad905 simplified printer context: uniform externing and static token translations;
wenzelm
parents: 42264
diff changeset
   663
  let
9566078ad905 simplified printer context: uniform externing and static token translations;
wenzelm
parents: 42264
diff changeset
   664
    val m =
9566078ad905 simplified printer context: uniform externing and static token translations;
wenzelm
parents: 42264
diff changeset
   665
      if Variable.is_fixed ctxt x orelse Syntax.is_pretty_global ctxt
50201
c26369c9eda6 Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents: 49690
diff changeset
   666
      then Markup.fixed x else Markup.intensify;
42267
9566078ad905 simplified printer context: uniform externing and static token translations;
wenzelm
parents: 42264
diff changeset
   667
  in
9566078ad905 simplified printer context: uniform externing and static token translations;
wenzelm
parents: 42264
diff changeset
   668
    if can Name.dest_skolem x
50201
c26369c9eda6 Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents: 49690
diff changeset
   669
    then ([m, Markup.skolem], Variable.revert_fixed ctxt x)
c26369c9eda6 Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents: 49690
diff changeset
   670
    else ([m, Markup.free], x)
42267
9566078ad905 simplified printer context: uniform externing and static token translations;
wenzelm
parents: 42264
diff changeset
   671
  end;
9566078ad905 simplified printer context: uniform externing and static token translations;
wenzelm
parents: 42264
diff changeset
   672
9566078ad905 simplified printer context: uniform externing and static token translations;
wenzelm
parents: 42264
diff changeset
   673
fun var_or_skolem s =
9566078ad905 simplified printer context: uniform externing and static token translations;
wenzelm
parents: 42264
diff changeset
   674
  (case Lexicon.read_variable s of
9566078ad905 simplified printer context: uniform externing and static token translations;
wenzelm
parents: 42264
diff changeset
   675
    SOME (x, i) =>
9566078ad905 simplified printer context: uniform externing and static token translations;
wenzelm
parents: 42264
diff changeset
   676
      (case try Name.dest_skolem x of
50201
c26369c9eda6 Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents: 49690
diff changeset
   677
        NONE => (Markup.var, s)
c26369c9eda6 Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents: 49690
diff changeset
   678
      | SOME x' => (Markup.skolem, Term.string_of_vname (x', i)))
c26369c9eda6 Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents: 49690
diff changeset
   679
  | NONE => (Markup.var, s));
42267
9566078ad905 simplified printer context: uniform externing and static token translations;
wenzelm
parents: 42264
diff changeset
   680
50201
c26369c9eda6 Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents: 49690
diff changeset
   681
val typing_elem = YXML.output_markup_elem Markup.typing;
c26369c9eda6 Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents: 49690
diff changeset
   682
val sorting_elem = YXML.output_markup_elem Markup.sorting;
49657
40e4feac2921 turn constraints into Isabelle_Markup.typing, depending on show_markup options;
wenzelm
parents: 49655
diff changeset
   683
42267
9566078ad905 simplified printer context: uniform externing and static token translations;
wenzelm
parents: 42264
diff changeset
   684
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
   685
  let
49686
678aa92e921c printed "_ofsort" is subject to show_markup as well;
wenzelm
parents: 49685
diff changeset
   686
    val show_markup = Config.get ctxt show_markup;
678aa92e921c printed "_ofsort" is subject to show_markup as well;
wenzelm
parents: 49685
diff changeset
   687
    val show_sorts = Config.get ctxt show_sorts;
52185
1b481b490454 discontinued obsolete show_all_types;
wenzelm
parents: 52177
diff changeset
   688
    val show_types = Config.get ctxt show_types orelse show_sorts;
49657
40e4feac2921 turn constraints into Isabelle_Markup.typing, depending on show_markup options;
wenzelm
parents: 49655
diff changeset
   689
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 42359
diff changeset
   690
    val syn = Proof_Context.syn_of ctxt;
49657
40e4feac2921 turn constraints into Isabelle_Markup.typing, depending on show_markup options;
wenzelm
parents: 49655
diff changeset
   691
    val prtabs = Syntax.prtabs syn;
40e4feac2921 turn constraints into Isabelle_Markup.typing, depending on show_markup options;
wenzelm
parents: 49655
diff changeset
   692
    val trf = Syntax.print_ast_translation syn;
40e4feac2921 turn constraints into Isabelle_Markup.typing, depending on show_markup options;
wenzelm
parents: 49655
diff changeset
   693
40e4feac2921 turn constraints into Isabelle_Markup.typing, depending on show_markup options;
wenzelm
parents: 49655
diff changeset
   694
    fun markup_extern c =
40e4feac2921 turn constraints into Isabelle_Markup.typing, depending on show_markup options;
wenzelm
parents: 49655
diff changeset
   695
      (case Syntax.lookup_const syn c of
40e4feac2921 turn constraints into Isabelle_Markup.typing, depending on show_markup options;
wenzelm
parents: 49655
diff changeset
   696
        SOME "" => ([], c)
40e4feac2921 turn constraints into Isabelle_Markup.typing, depending on show_markup options;
wenzelm
parents: 49655
diff changeset
   697
      | SOME b => markup_extern b
40e4feac2921 turn constraints into Isabelle_Markup.typing, depending on show_markup options;
wenzelm
parents: 49655
diff changeset
   698
      | NONE => c |> Lexicon.unmark
40e4feac2921 turn constraints into Isabelle_Markup.typing, depending on show_markup options;
wenzelm
parents: 49655
diff changeset
   699
         {case_class = fn x => (markup_class ctxt x, Proof_Context.extern_class ctxt x),
40e4feac2921 turn constraints into Isabelle_Markup.typing, depending on show_markup options;
wenzelm
parents: 49655
diff changeset
   700
          case_type = fn x => (markup_type ctxt x, Proof_Context.extern_type ctxt x),
40e4feac2921 turn constraints into Isabelle_Markup.typing, depending on show_markup options;
wenzelm
parents: 49655
diff changeset
   701
          case_const = fn x => (markup_const ctxt x, Proof_Context.extern_const ctxt x),
40e4feac2921 turn constraints into Isabelle_Markup.typing, depending on show_markup options;
wenzelm
parents: 49655
diff changeset
   702
          case_fixed = fn x => free_or_skolem ctxt x,
40e4feac2921 turn constraints into Isabelle_Markup.typing, depending on show_markup options;
wenzelm
parents: 49655
diff changeset
   703
          case_default = fn x => ([], x)});
42267
9566078ad905 simplified printer context: uniform externing and static token translations;
wenzelm
parents: 42264
diff changeset
   704
50201
c26369c9eda6 Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents: 49690
diff changeset
   705
    fun token_trans "_tfree" x = SOME (Pretty.mark_str (Markup.tfree, x))
c26369c9eda6 Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents: 49690
diff changeset
   706
      | token_trans "_tvar" x = SOME (Pretty.mark_str (Markup.tvar, x))
42267
9566078ad905 simplified printer context: uniform externing and static token translations;
wenzelm
parents: 42264
diff changeset
   707
      | token_trans "_free" x = SOME (Pretty.marks_str (free_or_skolem ctxt x))
50201
c26369c9eda6 Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents: 49690
diff changeset
   708
      | token_trans "_bound" x = SOME (Pretty.mark_str (Markup.bound, x))
c26369c9eda6 Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents: 49690
diff changeset
   709
      | token_trans "_loose" x = SOME (Pretty.mark_str (Markup.bad, x))
42267
9566078ad905 simplified printer context: uniform externing and static token translations;
wenzelm
parents: 42264
diff changeset
   710
      | token_trans "_var" x = SOME (Pretty.mark_str (var_or_skolem x))
50201
c26369c9eda6 Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents: 49690
diff changeset
   711
      | token_trans "_numeral" x = SOME (Pretty.mark_str (Markup.numeral, x))
c26369c9eda6 Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents: 49690
diff changeset
   712
      | token_trans "_inner_string" x = SOME (Pretty.mark_str (Markup.inner_string, x))
42267
9566078ad905 simplified printer context: uniform externing and static token translations;
wenzelm
parents: 42264
diff changeset
   713
      | token_trans _ _ = NONE;
9566078ad905 simplified printer context: uniform externing and static token translations;
wenzelm
parents: 42264
diff changeset
   714
49655
6642e559f165 tuned signature;
wenzelm
parents: 49358
diff changeset
   715
    fun markup_trans a [Ast.Variable x] = token_trans a x
49662
de6be6922c19 proper handling of constraints stemming from idtyp_ast_tr';
wenzelm
parents: 49660
diff changeset
   716
      | markup_trans "_constrain" [t, ty] = constrain_trans t ty
de6be6922c19 proper handling of constraints stemming from idtyp_ast_tr';
wenzelm
parents: 49660
diff changeset
   717
      | markup_trans "_idtyp" [t, ty] = constrain_trans t ty
49686
678aa92e921c printed "_ofsort" is subject to show_markup as well;
wenzelm
parents: 49685
diff changeset
   718
      | markup_trans "_ofsort" [ty, s] = ofsort_trans ty s
49657
40e4feac2921 turn constraints into Isabelle_Markup.typing, depending on show_markup options;
wenzelm
parents: 49655
diff changeset
   719
      | markup_trans _ _ = NONE
49655
6642e559f165 tuned signature;
wenzelm
parents: 49358
diff changeset
   720
49662
de6be6922c19 proper handling of constraints stemming from idtyp_ast_tr';
wenzelm
parents: 49660
diff changeset
   721
    and constrain_trans t ty =
49686
678aa92e921c printed "_ofsort" is subject to show_markup as well;
wenzelm
parents: 49685
diff changeset
   722
      if show_markup andalso not show_types then
49662
de6be6922c19 proper handling of constraints stemming from idtyp_ast_tr';
wenzelm
parents: 49660
diff changeset
   723
        let
de6be6922c19 proper handling of constraints stemming from idtyp_ast_tr';
wenzelm
parents: 49660
diff changeset
   724
          val ((bg1, bg2), en) = typing_elem;
de6be6922c19 proper handling of constraints stemming from idtyp_ast_tr';
wenzelm
parents: 49660
diff changeset
   725
          val bg = bg1 ^ Pretty.symbolic_output (pretty_typ_ast Markup.empty ty) ^ bg2;
de6be6922c19 proper handling of constraints stemming from idtyp_ast_tr';
wenzelm
parents: 49660
diff changeset
   726
        in SOME (Pretty.raw_markup (bg, en) (0, [pretty_ast Markup.empty t])) end
de6be6922c19 proper handling of constraints stemming from idtyp_ast_tr';
wenzelm
parents: 49660
diff changeset
   727
      else NONE
de6be6922c19 proper handling of constraints stemming from idtyp_ast_tr';
wenzelm
parents: 49660
diff changeset
   728
49686
678aa92e921c printed "_ofsort" is subject to show_markup as well;
wenzelm
parents: 49685
diff changeset
   729
    and ofsort_trans ty s =
678aa92e921c printed "_ofsort" is subject to show_markup as well;
wenzelm
parents: 49685
diff changeset
   730
      if show_markup andalso not show_sorts then
678aa92e921c printed "_ofsort" is subject to show_markup as well;
wenzelm
parents: 49685
diff changeset
   731
        let
678aa92e921c printed "_ofsort" is subject to show_markup as well;
wenzelm
parents: 49685
diff changeset
   732
          val ((bg1, bg2), en) = sorting_elem;
678aa92e921c printed "_ofsort" is subject to show_markup as well;
wenzelm
parents: 49685
diff changeset
   733
          val bg = bg1 ^ Pretty.symbolic_output (pretty_typ_ast Markup.empty s) ^ bg2;
678aa92e921c printed "_ofsort" is subject to show_markup as well;
wenzelm
parents: 49685
diff changeset
   734
        in SOME (Pretty.raw_markup (bg, en) (0, [pretty_typ_ast Markup.empty ty])) end
678aa92e921c printed "_ofsort" is subject to show_markup as well;
wenzelm
parents: 49685
diff changeset
   735
      else NONE
678aa92e921c printed "_ofsort" is subject to show_markup as well;
wenzelm
parents: 49685
diff changeset
   736
49657
40e4feac2921 turn constraints into Isabelle_Markup.typing, depending on show_markup options;
wenzelm
parents: 49655
diff changeset
   737
    and pretty_typ_ast m ast = ast
40e4feac2921 turn constraints into Isabelle_Markup.typing, depending on show_markup options;
wenzelm
parents: 49655
diff changeset
   738
      |> Printer.pretty_typ_ast ctxt prtabs trf markup_trans markup_extern
40e4feac2921 turn constraints into Isabelle_Markup.typing, depending on show_markup options;
wenzelm
parents: 49655
diff changeset
   739
      |> Pretty.markup m
40e4feac2921 turn constraints into Isabelle_Markup.typing, depending on show_markup options;
wenzelm
parents: 49655
diff changeset
   740
40e4feac2921 turn constraints into Isabelle_Markup.typing, depending on show_markup options;
wenzelm
parents: 49655
diff changeset
   741
    and pretty_ast m ast = ast
40e4feac2921 turn constraints into Isabelle_Markup.typing, depending on show_markup options;
wenzelm
parents: 49655
diff changeset
   742
      |> prt_t ctxt prtabs trf markup_trans markup_extern
40e4feac2921 turn constraints into Isabelle_Markup.typing, depending on show_markup options;
wenzelm
parents: 49655
diff changeset
   743
      |> Pretty.markup m;
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   744
  in
42255
097c93dcd877 eliminated slightly odd Syntax.rep_syntax;
wenzelm
parents: 42254
diff changeset
   745
    t_to_ast ctxt (Syntax.print_translation syn) t
097c93dcd877 eliminated slightly odd Syntax.rep_syntax;
wenzelm
parents: 42254
diff changeset
   746
    |> Ast.normalize ctxt (Syntax.print_rules syn)
49657
40e4feac2921 turn constraints into Isabelle_Markup.typing, depending on show_markup options;
wenzelm
parents: 49655
diff changeset
   747
    |> pretty_ast markup
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   748
  end;
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   749
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   750
in
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   751
50201
c26369c9eda6 Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents: 49690
diff changeset
   752
val unparse_sort = unparse_t sort_to_ast Printer.pretty_typ_ast Markup.sort;
c26369c9eda6 Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents: 49690
diff changeset
   753
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
   754
dd8029f71e1c separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff changeset
   755
fun unparse_term ctxt =
dd8029f71e1c separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff changeset
   756
  let
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 42359
diff changeset
   757
    val thy = Proof_Context.theory_of ctxt;
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 42359
diff changeset
   758
    val syn = Proof_Context.syn_of ctxt;
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 42359
diff changeset
   759
    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
   760
  in
42298
d622145603ee more accurate markup for syntax consts, notably binders which point back to the original logical entity;
wenzelm
parents: 42296
diff changeset
   761
    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
   762
      (Printer.pretty_term_ast (not (Pure_Thy.old_appl_syntax thy)))
50201
c26369c9eda6 Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents: 49690
diff changeset
   763
      Markup.term ctxt
42241
dd8029f71e1c separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff changeset
   764
  end;
dd8029f71e1c separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff changeset
   765
42249
12a073670584 simplified standard parse/unparse;
wenzelm
parents: 42248
diff changeset
   766
end;
12a073670584 simplified standard parse/unparse;
wenzelm
parents: 42248
diff changeset
   767
42241
dd8029f71e1c separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff changeset
   768
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   769
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   770
(** translations **)
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   771
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   772
(* type propositions *)
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   773
42247
12fe41a92cd5 typed_print_translation: discontinued show_sorts argument;
wenzelm
parents: 42245
diff changeset
   774
fun type_prop_tr' ctxt T [Const ("\\<^const>Pure.sort_constraint", _)] =
42476
d0bc1268ef09 clarified auxiliary structure Lexicon.Syntax;
wenzelm
parents: 42470
diff changeset
   775
      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
   776
  | type_prop_tr' ctxt T [t] =
42476
d0bc1268ef09 clarified auxiliary structure Lexicon.Syntax;
wenzelm
parents: 42470
diff changeset
   777
      Syntax.const "_ofclass" $ term_of_typ ctxt T $ t
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   778
  | 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
   779
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   780
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   781
(* type reflection *)
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   782
42247
12fe41a92cd5 typed_print_translation: discontinued show_sorts argument;
wenzelm
parents: 42245
diff changeset
   783
fun type_tr' ctxt (Type ("itself", [T])) ts =
42476
d0bc1268ef09 clarified auxiliary structure Lexicon.Syntax;
wenzelm
parents: 42470
diff changeset
   784
      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
   785
  | type_tr' _ _ _ = raise Match;
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   786
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   787
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   788
(* type constraints *)
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   789
42247
12fe41a92cd5 typed_print_translation: discontinued show_sorts argument;
wenzelm
parents: 42245
diff changeset
   790
fun type_constraint_tr' ctxt (Type ("fun", [T, _])) (t :: ts) =
42476
d0bc1268ef09 clarified auxiliary structure Lexicon.Syntax;
wenzelm
parents: 42470
diff changeset
   791
      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
   792
  | type_constraint_tr' _ _ _ = raise Match;
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   793
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   794
42295
8fdbb3b10beb moved CONST syntax/translations to their proper place;
wenzelm
parents: 42294
diff changeset
   795
(* authentic syntax *)
8fdbb3b10beb moved CONST syntax/translations to their proper place;
wenzelm
parents: 42294
diff changeset
   796
8fdbb3b10beb moved CONST syntax/translations to their proper place;
wenzelm
parents: 42294
diff changeset
   797
fun const_ast_tr intern ctxt [Ast.Variable c] =
8fdbb3b10beb moved CONST syntax/translations to their proper place;
wenzelm
parents: 42294
diff changeset
   798
      let
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 42359
diff changeset
   799
        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
   800
        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
   801
      in Ast.Constant d end
42470
cc78b0ed0fad more precise error positions;
wenzelm
parents: 42410
diff changeset
   802
  | 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
   803
      (Ast.Appl [Ast.Constant "_constrain", const_ast_tr intern ctxt [x], T]
cc78b0ed0fad more precise error positions;
wenzelm
parents: 42410
diff changeset
   804
        handle ERROR msg =>
48992
0518bf89c777 renamed Position.str_of to Position.here;
wenzelm
parents: 48768
diff changeset
   805
          error (msg ^ Position.here (the_default Position.none (Term_Position.decode pos))))
42295
8fdbb3b10beb moved CONST syntax/translations to their proper place;
wenzelm
parents: 42294
diff changeset
   806
  | 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
   807
8fdbb3b10beb moved CONST syntax/translations to their proper place;
wenzelm
parents: 42294
diff changeset
   808
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   809
(* setup translations *)
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   810
53171
a5e54d4d9081 added Theory.setup convenience;
wenzelm
parents: 52517
diff changeset
   811
val _ = Theory.setup
52143
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 50201
diff changeset
   812
 (Sign.parse_ast_translation
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 50201
diff changeset
   813
   [("_context_const", const_ast_tr true),
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 50201
diff changeset
   814
    ("_context_xconst", const_ast_tr false)] #>
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 50201
diff changeset
   815
  Sign.typed_print_translation
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   816
   [("_type_prop", type_prop_tr'),
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   817
    ("\\<^const>TYPE", type_tr'),
53171
a5e54d4d9081 added Theory.setup convenience;
wenzelm
parents: 52517
diff changeset
   818
    ("_type_constraint_", type_constraint_tr')]);
42245
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   819
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   820
29e3967550d5 moved unparse material to syntax_phases.ML;
wenzelm
parents: 42243
diff changeset
   821
42382
dcd983ee2c29 provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents: 42379
diff changeset
   822
(** check/uncheck **)
dcd983ee2c29 provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents: 42379
diff changeset
   823
45429
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   824
(* context-sensitive (un)checking *)
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   825
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   826
type key = int * bool;
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   827
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   828
structure Checks = Generic_Data
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   829
(
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   830
  type 'a check = 'a list -> Proof.context -> ('a list * Proof.context) option;
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   831
  type T =
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   832
    ((key * ((string * typ check) * stamp) list) list *
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   833
     (key * ((string * term check) * stamp) list) list);
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   834
  val empty = ([], []);
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   835
  val extend = I;
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   836
  fun merge ((typ_checks1, term_checks1), (typ_checks2, term_checks2)) : T =
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   837
    (AList.join (op =) (K (Library.merge (eq_snd (op =)))) (typ_checks1, typ_checks2),
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   838
     AList.join (op =) (K (Library.merge (eq_snd (op =)))) (term_checks1, term_checks2));
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   839
);
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   840
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   841
fun print_checks ctxt =
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   842
  let
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   843
    fun split_checks checks =
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   844
      List.partition (fn ((_, un), _) => not un) checks
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   845
      |> pairself (map (fn ((i, _), fs) => (i, map (fst o fst) fs))
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   846
          #> sort (int_ord o pairself fst));
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   847
    fun pretty_checks kind checks =
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   848
      checks |> map (fn (i, names) => Pretty.block
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   849
        [Pretty.str (kind ^ " (stage " ^ signed_string_of_int i ^ "):"),
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   850
          Pretty.brk 1, Pretty.strs names]);
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   851
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   852
    val (typs, terms) = Checks.get (Context.Proof ctxt);
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   853
    val (typ_checks, typ_unchecks) = split_checks typs;
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   854
    val (term_checks, term_unchecks) = split_checks terms;
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   855
  in
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   856
    pretty_checks "typ_checks" typ_checks @
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   857
    pretty_checks "term_checks" term_checks @
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   858
    pretty_checks "typ_unchecks" typ_unchecks @
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   859
    pretty_checks "term_unchecks" term_unchecks
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   860
  end |> Pretty.chunks |> Pretty.writeln;
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   861
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   862
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   863
local
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   864
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   865
fun context_check which (key: key) name f =
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   866
  Checks.map (which (AList.map_default op = (key, []) (cons ((name, f), stamp ()))));
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   867
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   868
fun simple_check eq f xs ctxt =
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   869
  let val xs' = f ctxt xs
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   870
  in if eq_list eq (xs, xs') then NONE else SOME (xs', ctxt) end;
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   871
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   872
in
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   873
45444
ac069060e08a tuned signature;
wenzelm
parents: 45429
diff changeset
   874
fun typ_check' stage = context_check apfst (stage, false);
ac069060e08a tuned signature;
wenzelm
parents: 45429
diff changeset
   875
fun term_check' stage = context_check apsnd (stage, false);
ac069060e08a tuned signature;
wenzelm
parents: 45429
diff changeset
   876
fun typ_uncheck' stage = context_check apfst (stage, true);
ac069060e08a tuned signature;
wenzelm
parents: 45429
diff changeset
   877
fun term_uncheck' stage = context_check apsnd (stage, true);
45429
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   878
45444
ac069060e08a tuned signature;
wenzelm
parents: 45429
diff changeset
   879
fun typ_check key name f = typ_check' key name (simple_check (op =) f);
ac069060e08a tuned signature;
wenzelm
parents: 45429
diff changeset
   880
fun term_check key name f = term_check' key name (simple_check (op aconv) f);
ac069060e08a tuned signature;
wenzelm
parents: 45429
diff changeset
   881
fun typ_uncheck key name f = typ_uncheck' key name (simple_check (op =) f);
ac069060e08a tuned signature;
wenzelm
parents: 45429
diff changeset
   882
fun term_uncheck key name f = term_uncheck' key name (simple_check (op aconv) f);
45429
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   883
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   884
end;
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   885
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   886
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   887
local
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   888
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   889
fun check_stage fs = perhaps_loop (perhaps_apply (map uncurry fs));
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   890
fun check_all fs = perhaps_apply (map check_stage fs);
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   891
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   892
fun check which uncheck ctxt0 xs0 =
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   893
  let
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   894
    val funs = which (Checks.get (Context.Proof ctxt0))
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   895
      |> 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
   896
      |> Library.sort (int_ord o pairself fst) |> map snd
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   897
      |> not uncheck ? map rev;
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   898
  in #1 (perhaps (check_all funs) (xs0, ctxt0)) end;
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   899
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   900
val apply_typ_check = check fst false;
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   901
val apply_term_check = check snd false;
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   902
val apply_typ_uncheck = check fst true;
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   903
val apply_term_uncheck = check snd true;
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   904
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   905
in
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   906
49674
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 49662
diff changeset
   907
fun check_typs ctxt raw_tys =
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 49662
diff changeset
   908
  let
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 49662
diff changeset
   909
    val (sorting_report, tys) = Proof_Context.prepare_sortsT ctxt raw_tys;
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 49662
diff changeset
   910
    val _ = Context_Position.if_visible ctxt Output.report sorting_report;
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 49662
diff changeset
   911
  in
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 49662
diff changeset
   912
    tys
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 49662
diff changeset
   913
    |> apply_typ_check ctxt
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 49662
diff changeset
   914
    |> Term_Sharing.typs (Proof_Context.theory_of ctxt)
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 49662
diff changeset
   915
  end;
45423
92f91f990165 tuned signature -- emphasize internal role of these operations;
wenzelm
parents: 45412
diff changeset
   916
45445
41e641a870de pass term positions into check phase, where resulting types are reported accordingly, and eventually shown as tooltips;
wenzelm
parents: 45444
diff changeset
   917
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
   918
  let
49674
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 49662
diff changeset
   919
    val (sorting_report, raw_ts') = Proof_Context.prepare_sorts ctxt raw_ts;
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 49662
diff changeset
   920
    val (ts, ps) = Type_Infer_Context.prepare_positions ctxt raw_ts';
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 49662
diff changeset
   921
45445
41e641a870de pass term positions into check phase, where resulting types are reported accordingly, and eventually shown as tooltips;
wenzelm
parents: 45444
diff changeset
   922
    val tys = map (Logic.mk_type o snd) ps;
45448
018f8959c7a6 more efficient prepare_sorts -- bypass encoded positions;
wenzelm
parents: 45447
diff changeset
   923
    val (ts', tys') = ts @ tys
45445
41e641a870de pass term positions into check phase, where resulting types are reported accordingly, and eventually shown as tooltips;
wenzelm
parents: 45444
diff changeset
   924
      |> 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
   925
      |> chop (length ts);
49674
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 49662
diff changeset
   926
    val typing_report =
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 49662
diff changeset
   927
      fold2 (fn (pos, _) => fn ty =>
45445
41e641a870de pass term positions into check phase, where resulting types are reported accordingly, and eventually shown as tooltips;
wenzelm
parents: 45444
diff changeset
   928
        if Position.is_reported pos then
50201
c26369c9eda6 Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents: 49690
diff changeset
   929
          cons (Position.reported_text pos Markup.typing
49674
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 49662
diff changeset
   930
            (Syntax.string_of_typ ctxt (Logic.dest_type ty)))
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 49662
diff changeset
   931
        else I) ps tys' []
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 49662
diff changeset
   932
      |> implode;
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 49662
diff changeset
   933
dbadb4d03cbc report sort assignment of visible type variables;
wenzelm
parents: 49662
diff changeset
   934
    val _ = Context_Position.if_visible ctxt Output.report (sorting_report ^ typing_report);
45445
41e641a870de pass term positions into check phase, where resulting types are reported accordingly, and eventually shown as tooltips;
wenzelm
parents: 45444
diff changeset
   935
  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
   936
42382
dcd983ee2c29 provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents: 42379
diff changeset
   937
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
   938
45429
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   939
val uncheck_typs = apply_typ_uncheck;
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   940
val uncheck_terms = apply_term_uncheck;
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   941
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   942
end;
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   943
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   944
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   945
(* standard phases *)
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   946
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   947
val _ = Context.>>
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   948
 (typ_check 0 "standard" Proof_Context.standard_typ_check #>
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   949
  term_check 0 "standard"
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   950
    (fn ctxt => Type_Infer_Context.infer_types ctxt #> map (Proof_Context.expand_abbrevs ctxt)) #>
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   951
  term_check 100 "standard_finish" Proof_Context.standard_term_check_finish #>
fd58cbf8cae3 tuned signature;
wenzelm
parents: 45427
diff changeset
   952
  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
   953
dcd983ee2c29 provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents: 42379
diff changeset
   954
dcd983ee2c29 provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents: 42379
diff changeset
   955
42242
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   956
(** install operations **)
39261908e12f moved decode/parse operations to standard_syntax.ML;
wenzelm
parents: 42241
diff changeset
   957
42241
dd8029f71e1c separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff changeset
   958
val _ = Syntax.install_operations
dd8029f71e1c separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff changeset
   959
  {parse_sort = parse_sort,
dd8029f71e1c separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff changeset
   960
   parse_typ = parse_typ,
42281
69d4543811d0 simplified read_term vs. read_prop;
wenzelm
parents: 42267
diff changeset
   961
   parse_term = parse_term false,
69d4543811d0 simplified read_term vs. read_prop;
wenzelm
parents: 42267
diff changeset
   962
   parse_prop = parse_term true,
42241
dd8029f71e1c separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff changeset
   963
   unparse_sort = unparse_sort,
dd8029f71e1c separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff changeset
   964
   unparse_typ = unparse_typ,
42382
dcd983ee2c29 provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents: 42379
diff changeset
   965
   unparse_term = unparse_term,
dcd983ee2c29 provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents: 42379
diff changeset
   966
   check_typs = check_typs,
dcd983ee2c29 provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents: 42379
diff changeset
   967
   check_terms = check_terms,
dcd983ee2c29 provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents: 42379
diff changeset
   968
   check_props = check_props,
dcd983ee2c29 provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents: 42379
diff changeset
   969
   uncheck_typs = uncheck_typs,
dcd983ee2c29 provide structure Syntax early (before structure Type), back-patch check/uncheck later;
wenzelm
parents: 42379
diff changeset
   970
   uncheck_terms = uncheck_terms};
42241
dd8029f71e1c separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff changeset
   971
dd8029f71e1c separate module for standard implementation of inner syntax operations;
wenzelm
parents:
diff changeset
   972
end;