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