src/Pure/Thy/thy_syntax.ML
author wenzelm
Tue, 12 Aug 2014 00:08:32 +0200
changeset 57905 c0c5652e796e
parent 57903 ade8d65b212e
child 58864 505a8150368a
permissions -rw-r--r--
separate module Command_Span: mostly syntactic representation; potentially prover-specific Output_Syntax.parse_spans;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
29315
b074c05f00ad renamed ThyEdit (in thy_edit.ML) to ThySyntax (in thy_syntax.ML);
wenzelm
parents: 28454
diff changeset
     1
(*  Title:      Pure/Thy/thy_syntax.ML
23726
2ebecff57b17 Basic editing of theory sources.
wenzelm
parents:
diff changeset
     2
    Author:     Makarius
2ebecff57b17 Basic editing of theory sources.
wenzelm
parents:
diff changeset
     3
29315
b074c05f00ad renamed ThyEdit (in thy_edit.ML) to ThySyntax (in thy_syntax.ML);
wenzelm
parents: 28454
diff changeset
     4
Superficial theory syntax: tokens and spans.
23726
2ebecff57b17 Basic editing of theory sources.
wenzelm
parents:
diff changeset
     5
*)
2ebecff57b17 Basic editing of theory sources.
wenzelm
parents:
diff changeset
     6
29315
b074c05f00ad renamed ThyEdit (in thy_edit.ML) to ThySyntax (in thy_syntax.ML);
wenzelm
parents: 28454
diff changeset
     7
signature THY_SYNTAX =
23726
2ebecff57b17 Basic editing of theory sources.
wenzelm
parents:
diff changeset
     8
sig
57105
bf5ddf4ec64b tuned signature;
wenzelm
parents: 55915
diff changeset
     9
  val reports_of_tokens: Token.T list -> bool * Position.report_text list
40131
7cbebd636e79 explicitly qualify type Output.output, which is a slightly odd internal feature;
wenzelm
parents: 39507
diff changeset
    10
  val present_token: Token.T -> Output.output
57905
c0c5652e796e separate module Command_Span: mostly syntactic representation;
wenzelm
parents: 57903
diff changeset
    11
  val present_span: Command_Span.span -> Output.output
51225
3fe0d8d55975 support nested Thy_Syntax.element;
wenzelm
parents: 50238
diff changeset
    12
  datatype 'a element = Element of 'a * ('a element list * 'a) option
51321
75682dfff630 more Thy_Syntax.element operations;
wenzelm
parents: 51268
diff changeset
    13
  val atom: 'a -> 'a element
51225
3fe0d8d55975 support nested Thy_Syntax.element;
wenzelm
parents: 50238
diff changeset
    14
  val map_element: ('a -> 'b) -> 'a element -> 'b element
3fe0d8d55975 support nested Thy_Syntax.element;
wenzelm
parents: 50238
diff changeset
    15
  val flat_element: 'a element -> 'a list
51268
fcc4b89a600d simplified Outer_Syntax.read_span: internalized Toplevel.is_ignored;
wenzelm
parents: 51267
diff changeset
    16
  val last_element: 'a element -> 'a
57905
c0c5652e796e separate module Command_Span: mostly syntactic representation;
wenzelm
parents: 57903
diff changeset
    17
  val parse_elements: Command_Span.span list -> Command_Span.span element list
23726
2ebecff57b17 Basic editing of theory sources.
wenzelm
parents:
diff changeset
    18
end;
2ebecff57b17 Basic editing of theory sources.
wenzelm
parents:
diff changeset
    19
37216
3165bc303f66 modernized some structure names, keeping a few legacy aliases;
wenzelm
parents: 37197
diff changeset
    20
structure Thy_Syntax: THY_SYNTAX =
23726
2ebecff57b17 Basic editing of theory sources.
wenzelm
parents:
diff changeset
    21
struct
2ebecff57b17 Basic editing of theory sources.
wenzelm
parents:
diff changeset
    22
57905
c0c5652e796e separate module Command_Span: mostly syntactic representation;
wenzelm
parents: 57903
diff changeset
    23
(** presentation **)
23726
2ebecff57b17 Basic editing of theory sources.
wenzelm
parents:
diff changeset
    24
2ebecff57b17 Basic editing of theory sources.
wenzelm
parents:
diff changeset
    25
local
2ebecff57b17 Basic editing of theory sources.
wenzelm
parents:
diff changeset
    26
48749
c197b3c3e7fa some attempts to keep malformed syntax errors focussed, without too much red spilled onto the document view;
wenzelm
parents: 46811
diff changeset
    27
fun reports_of_token tok =
c197b3c3e7fa some attempts to keep malformed syntax errors focussed, without too much red spilled onto the document view;
wenzelm
parents: 46811
diff changeset
    28
  let
55828
42ac3cfb89f6 clarified language markup: added "delimited" property;
wenzelm
parents: 55744
diff changeset
    29
    val {text, pos, ...} = Token.source_position_of tok;
48749
c197b3c3e7fa some attempts to keep malformed syntax errors focussed, without too much red spilled onto the document view;
wenzelm
parents: 46811
diff changeset
    30
    val malformed_symbols =
55828
42ac3cfb89f6 clarified language markup: added "delimited" property;
wenzelm
parents: 55744
diff changeset
    31
      Symbol_Pos.explode (text, pos)
48749
c197b3c3e7fa some attempts to keep malformed syntax errors focussed, without too much red spilled onto the document view;
wenzelm
parents: 46811
diff changeset
    32
      |> map_filter (fn (sym, pos) =>
48756
1c843142758e sneak message into "bad" markup as property -- to be displayed after YXML parsing;
wenzelm
parents: 48752
diff changeset
    33
          if Symbol.is_malformed sym
50201
c26369c9eda6 Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
wenzelm
parents: 48878
diff changeset
    34
          then SOME ((pos, Markup.bad), "Malformed symbolic character") else NONE);
48749
c197b3c3e7fa some attempts to keep malformed syntax errors focussed, without too much red spilled onto the document view;
wenzelm
parents: 46811
diff changeset
    35
    val is_malformed = Token.is_error tok orelse not (null malformed_symbols);
55915
607948c90bf0 suppress short abbreviations more uniformly, for outer and quasi-outer syntax;
wenzelm
parents: 55828
diff changeset
    36
    val reports = Token.report tok :: Token.completion_report tok @ malformed_symbols;
48749
c197b3c3e7fa some attempts to keep malformed syntax errors focussed, without too much red spilled onto the document view;
wenzelm
parents: 46811
diff changeset
    37
  in (is_malformed, reports) end;
40528
d5e9f7608341 report malformed symbols;
wenzelm
parents: 40523
diff changeset
    38
23803
11bf7af10ec8 tuned signature;
wenzelm
parents: 23790
diff changeset
    39
in
11bf7af10ec8 tuned signature;
wenzelm
parents: 23790
diff changeset
    40
48749
c197b3c3e7fa some attempts to keep malformed syntax errors focussed, without too much red spilled onto the document view;
wenzelm
parents: 46811
diff changeset
    41
fun reports_of_tokens toks =
c197b3c3e7fa some attempts to keep malformed syntax errors focussed, without too much red spilled onto the document view;
wenzelm
parents: 46811
diff changeset
    42
  let val results = map reports_of_token toks
c197b3c3e7fa some attempts to keep malformed syntax errors focussed, without too much red spilled onto the document view;
wenzelm
parents: 46811
diff changeset
    43
  in (exists fst results, maps snd results) end;
c197b3c3e7fa some attempts to keep malformed syntax errors focussed, without too much red spilled onto the document view;
wenzelm
parents: 46811
diff changeset
    44
57905
c0c5652e796e separate module Command_Span: mostly syntactic representation;
wenzelm
parents: 57903
diff changeset
    45
end;
c0c5652e796e separate module Command_Span: mostly syntactic representation;
wenzelm
parents: 57903
diff changeset
    46
23726
2ebecff57b17 Basic editing of theory sources.
wenzelm
parents:
diff changeset
    47
fun present_token tok =
55915
607948c90bf0 suppress short abbreviations more uniformly, for outer and quasi-outer syntax;
wenzelm
parents: 55828
diff changeset
    48
  Markup.enclose (Token.markup tok) (Output.output (Token.unparse tok));
23726
2ebecff57b17 Basic editing of theory sources.
wenzelm
parents:
diff changeset
    49
57905
c0c5652e796e separate module Command_Span: mostly syntactic representation;
wenzelm
parents: 57903
diff changeset
    50
val present_span = implode o map present_token o Command_Span.content;
54451
459cf6ee254e tuned signature;
wenzelm
parents: 53843
diff changeset
    51
459cf6ee254e tuned signature;
wenzelm
parents: 53843
diff changeset
    52
28434
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
    53
43621
e8858190cccd clarified Thy_Syntax.element;
wenzelm
parents: 43430
diff changeset
    54
(** specification elements: commands with optional proof **)
e8858190cccd clarified Thy_Syntax.element;
wenzelm
parents: 43430
diff changeset
    55
51225
3fe0d8d55975 support nested Thy_Syntax.element;
wenzelm
parents: 50238
diff changeset
    56
datatype 'a element = Element of 'a * ('a element list * 'a) option;
3fe0d8d55975 support nested Thy_Syntax.element;
wenzelm
parents: 50238
diff changeset
    57
3fe0d8d55975 support nested Thy_Syntax.element;
wenzelm
parents: 50238
diff changeset
    58
fun element (a, b) = Element (a, SOME b);
3fe0d8d55975 support nested Thy_Syntax.element;
wenzelm
parents: 50238
diff changeset
    59
fun atom a = Element (a, NONE);
43621
e8858190cccd clarified Thy_Syntax.element;
wenzelm
parents: 43430
diff changeset
    60
51225
3fe0d8d55975 support nested Thy_Syntax.element;
wenzelm
parents: 50238
diff changeset
    61
fun map_element f (Element (a, NONE)) = Element (f a, NONE)
3fe0d8d55975 support nested Thy_Syntax.element;
wenzelm
parents: 50238
diff changeset
    62
  | map_element f (Element (a, SOME (elems, b))) =
3fe0d8d55975 support nested Thy_Syntax.element;
wenzelm
parents: 50238
diff changeset
    63
      Element (f a, SOME ((map o map_element) f elems, f b));
3fe0d8d55975 support nested Thy_Syntax.element;
wenzelm
parents: 50238
diff changeset
    64
3fe0d8d55975 support nested Thy_Syntax.element;
wenzelm
parents: 50238
diff changeset
    65
fun flat_element (Element (a, NONE)) = [a]
3fe0d8d55975 support nested Thy_Syntax.element;
wenzelm
parents: 50238
diff changeset
    66
  | flat_element (Element (a, SOME (elems, b))) = a :: maps flat_element elems @ [b];
43621
e8858190cccd clarified Thy_Syntax.element;
wenzelm
parents: 43430
diff changeset
    67
51268
fcc4b89a600d simplified Outer_Syntax.read_span: internalized Toplevel.is_ignored;
wenzelm
parents: 51267
diff changeset
    68
fun last_element (Element (a, NONE)) = a
fcc4b89a600d simplified Outer_Syntax.read_span: internalized Toplevel.is_ignored;
wenzelm
parents: 51267
diff changeset
    69
  | last_element (Element (_, SOME (_, b))) = b;
fcc4b89a600d simplified Outer_Syntax.read_span: internalized Toplevel.is_ignored;
wenzelm
parents: 51267
diff changeset
    70
28434
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
    71
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
    72
(* scanning spans *)
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
    73
57905
c0c5652e796e separate module Command_Span: mostly syntactic representation;
wenzelm
parents: 57903
diff changeset
    74
val eof = Command_Span.Span (Command_Span.Command_Span ("", Position.none), []);
28434
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
    75
57905
c0c5652e796e separate module Command_Span: mostly syntactic representation;
wenzelm
parents: 57903
diff changeset
    76
fun is_eof (Command_Span.Span (Command_Span.Command_Span ("", _), _)) = true
28434
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
    77
  | is_eof _ = false;
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
    78
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
    79
val not_eof = not o is_eof;
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
    80
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
    81
val stopper = Scan.stopper (K eof) is_eof;
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
    82
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
    83
46811
03a2dc9e0624 clarified command span: include trailing whitespace/comments and thus reduce number of ignored spans with associated transactions and states (factor 2);
wenzelm
parents: 45666
diff changeset
    84
(* parse *)
28434
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
    85
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
    86
local
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
    87
48878
5e850e6fa3c3 tuned errors;
wenzelm
parents: 48867
diff changeset
    88
fun command_with pred =
57905
c0c5652e796e separate module Command_Span: mostly syntactic representation;
wenzelm
parents: 57903
diff changeset
    89
  Scan.one
c0c5652e796e separate module Command_Span: mostly syntactic representation;
wenzelm
parents: 57903
diff changeset
    90
    (fn (Command_Span.Span (Command_Span.Command_Span (name, _), _)) => pred name | _ => false);
28434
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
    91
51225
3fe0d8d55975 support nested Thy_Syntax.element;
wenzelm
parents: 50238
diff changeset
    92
val proof_atom =
57905
c0c5652e796e separate module Command_Span: mostly syntactic representation;
wenzelm
parents: 57903
diff changeset
    93
  Scan.one
c0c5652e796e separate module Command_Span: mostly syntactic representation;
wenzelm
parents: 57903
diff changeset
    94
    (fn (Command_Span.Span (Command_Span.Command_Span (name, _), _)) => Keyword.is_proof_body name
c0c5652e796e separate module Command_Span: mostly syntactic representation;
wenzelm
parents: 57903
diff changeset
    95
      | _ => true) >> atom;
28434
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
    96
51225
3fe0d8d55975 support nested Thy_Syntax.element;
wenzelm
parents: 50238
diff changeset
    97
fun proof_element x = (command_with Keyword.is_proof_goal -- proof_rest >> element || proof_atom) x
3fe0d8d55975 support nested Thy_Syntax.element;
wenzelm
parents: 50238
diff changeset
    98
and proof_rest x = (Scan.repeat proof_element -- command_with Keyword.is_qed) x;
3fe0d8d55975 support nested Thy_Syntax.element;
wenzelm
parents: 50238
diff changeset
    99
3fe0d8d55975 support nested Thy_Syntax.element;
wenzelm
parents: 50238
diff changeset
   100
val other_element =
3fe0d8d55975 support nested Thy_Syntax.element;
wenzelm
parents: 50238
diff changeset
   101
  command_with Keyword.is_theory_goal -- proof_rest >> element ||
3fe0d8d55975 support nested Thy_Syntax.element;
wenzelm
parents: 50238
diff changeset
   102
  Scan.one not_eof >> atom;
28434
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
   103
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
   104
in
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
   105
46811
03a2dc9e0624 clarified command span: include trailing whitespace/comments and thus reduce number of ignored spans with associated transactions and states (factor 2);
wenzelm
parents: 45666
diff changeset
   106
val parse_elements =
03a2dc9e0624 clarified command span: include trailing whitespace/comments and thus reduce number of ignored spans with associated transactions and states (factor 2);
wenzelm
parents: 45666
diff changeset
   107
  Source.of_list #>
51225
3fe0d8d55975 support nested Thy_Syntax.element;
wenzelm
parents: 50238
diff changeset
   108
  Source.source stopper (Scan.bulk other_element) NONE #>
46811
03a2dc9e0624 clarified command span: include trailing whitespace/comments and thus reduce number of ignored spans with associated transactions and states (factor 2);
wenzelm
parents: 45666
diff changeset
   109
  Source.exhaust;
28434
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
   110
23726
2ebecff57b17 Basic editing of theory sources.
wenzelm
parents:
diff changeset
   111
end;
28434
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
   112
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
   113
end;