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