src/Pure/Thy/thy_syntax.ML
author wenzelm
Fri, 10 Aug 2012 10:23:54 +0200
changeset 48752 8a81ef0bc790
parent 48751 dc3bbdda4bc8
child 48756 1c843142758e
permissions -rw-r--r--
discontinued mostly unused markup for command 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
36959
f5417836dbea renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
wenzelm
parents: 36950
diff changeset
     9
  val parse_tokens: Scan.lexicon * Scan.lexicon -> Position.T -> string -> Token.T list
48749
c197b3c3e7fa some attempts to keep malformed syntax errors focussed, without too much red spilled onto the document view;
wenzelm
parents: 46811
diff changeset
    10
  val reports_of_tokens: Token.T list -> bool * Position.report list
40131
7cbebd636e79 explicitly qualify type Output.output, which is a slightly odd internal feature;
wenzelm
parents: 39507
diff changeset
    11
  val present_token: Token.T -> Output.output
27842
6c35d50d309f abstract type span, tuned interfaces;
wenzelm
parents: 27770
diff changeset
    12
  datatype span_kind = Command of string | Ignored | Malformed
6c35d50d309f abstract type span, tuned interfaces;
wenzelm
parents: 27770
diff changeset
    13
  type span
6c35d50d309f abstract type span, tuned interfaces;
wenzelm
parents: 27770
diff changeset
    14
  val span_kind: span -> span_kind
36959
f5417836dbea renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
wenzelm
parents: 36950
diff changeset
    15
  val span_content: span -> Token.T list
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
    16
  val parse_spans: Token.T list -> span list
40131
7cbebd636e79 explicitly qualify type Output.output, which is a slightly odd internal feature;
wenzelm
parents: 39507
diff changeset
    17
  val present_span: span -> Output.output
43621
e8858190cccd clarified Thy_Syntax.element;
wenzelm
parents: 43430
diff changeset
    18
  type element = {head: span, proof: span list, proper_proof: bool}
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
    19
  val parse_elements: span list -> element list
23726
2ebecff57b17 Basic editing of theory sources.
wenzelm
parents:
diff changeset
    20
end;
2ebecff57b17 Basic editing of theory sources.
wenzelm
parents:
diff changeset
    21
37216
3165bc303f66 modernized some structure names, keeping a few legacy aliases;
wenzelm
parents: 37197
diff changeset
    22
structure Thy_Syntax: THY_SYNTAX =
23726
2ebecff57b17 Basic editing of theory sources.
wenzelm
parents:
diff changeset
    23
struct
2ebecff57b17 Basic editing of theory sources.
wenzelm
parents:
diff changeset
    24
23803
11bf7af10ec8 tuned signature;
wenzelm
parents: 23790
diff changeset
    25
(** tokens **)
11bf7af10ec8 tuned signature;
wenzelm
parents: 23790
diff changeset
    26
11bf7af10ec8 tuned signature;
wenzelm
parents: 23790
diff changeset
    27
(* parse *)
23726
2ebecff57b17 Basic editing of theory sources.
wenzelm
parents:
diff changeset
    28
40523
1050315f6ee2 simplified/robustified treatment of malformed symbols, which are now fully internalized (total Symbol.explode etc.);
wenzelm
parents: 40290
diff changeset
    29
fun parse_tokens lexs pos =
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
    30
  Source.of_string #>
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
    31
  Symbol.source #>
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
    32
  Token.source {do_recover = SOME false} (K lexs) pos #>
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
    33
  Source.exhaust;
23726
2ebecff57b17 Basic editing of theory sources.
wenzelm
parents:
diff changeset
    34
2ebecff57b17 Basic editing of theory sources.
wenzelm
parents:
diff changeset
    35
23803
11bf7af10ec8 tuned signature;
wenzelm
parents: 23790
diff changeset
    36
(* present *)
23726
2ebecff57b17 Basic editing of theory sources.
wenzelm
parents:
diff changeset
    37
2ebecff57b17 Basic editing of theory sources.
wenzelm
parents:
diff changeset
    38
local
2ebecff57b17 Basic editing of theory sources.
wenzelm
parents:
diff changeset
    39
2ebecff57b17 Basic editing of theory sources.
wenzelm
parents:
diff changeset
    40
val token_kind_markup =
45666
d83797ef0d2d separate module for concrete Isabelle markup;
wenzelm
parents: 44736
diff changeset
    41
 fn Token.Command       => Isabelle_Markup.command
d83797ef0d2d separate module for concrete Isabelle markup;
wenzelm
parents: 44736
diff changeset
    42
  | Token.Keyword       => Isabelle_Markup.keyword
44706
fe319b45315c eliminated markup for plain identifiers (frequent but insignificant);
wenzelm
parents: 44658
diff changeset
    43
  | Token.Ident         => Markup.empty
fe319b45315c eliminated markup for plain identifiers (frequent but insignificant);
wenzelm
parents: 44658
diff changeset
    44
  | Token.LongIdent     => Markup.empty
fe319b45315c eliminated markup for plain identifiers (frequent but insignificant);
wenzelm
parents: 44658
diff changeset
    45
  | Token.SymIdent      => Markup.empty
45666
d83797ef0d2d separate module for concrete Isabelle markup;
wenzelm
parents: 44736
diff changeset
    46
  | Token.Var           => Isabelle_Markup.var
d83797ef0d2d separate module for concrete Isabelle markup;
wenzelm
parents: 44736
diff changeset
    47
  | Token.TypeIdent     => Isabelle_Markup.tfree
d83797ef0d2d separate module for concrete Isabelle markup;
wenzelm
parents: 44736
diff changeset
    48
  | Token.TypeVar       => Isabelle_Markup.tvar
44706
fe319b45315c eliminated markup for plain identifiers (frequent but insignificant);
wenzelm
parents: 44658
diff changeset
    49
  | Token.Nat           => Markup.empty
fe319b45315c eliminated markup for plain identifiers (frequent but insignificant);
wenzelm
parents: 44658
diff changeset
    50
  | Token.Float         => Markup.empty
45666
d83797ef0d2d separate module for concrete Isabelle markup;
wenzelm
parents: 44736
diff changeset
    51
  | Token.String        => Isabelle_Markup.string
d83797ef0d2d separate module for concrete Isabelle markup;
wenzelm
parents: 44736
diff changeset
    52
  | Token.AltString     => Isabelle_Markup.altstring
d83797ef0d2d separate module for concrete Isabelle markup;
wenzelm
parents: 44736
diff changeset
    53
  | Token.Verbatim      => Isabelle_Markup.verbatim
38474
e498dc2eb576 uniform Markup.empty/Markup.Empty in ML and Scala;
wenzelm
parents: 38471
diff changeset
    54
  | Token.Space         => Markup.empty
45666
d83797ef0d2d separate module for concrete Isabelle markup;
wenzelm
parents: 44736
diff changeset
    55
  | Token.Comment       => Isabelle_Markup.comment
38474
e498dc2eb576 uniform Markup.empty/Markup.Empty in ML and Scala;
wenzelm
parents: 38471
diff changeset
    56
  | Token.InternalValue => Markup.empty
48751
dc3bbdda4bc8 more visible markup of malformed input as "bad";
wenzelm
parents: 48749
diff changeset
    57
  | Token.Error _       => Isabelle_Markup.bad
45666
d83797ef0d2d separate module for concrete Isabelle markup;
wenzelm
parents: 44736
diff changeset
    58
  | Token.Sync          => Isabelle_Markup.control
d83797ef0d2d separate module for concrete Isabelle markup;
wenzelm
parents: 44736
diff changeset
    59
  | Token.EOF           => Isabelle_Markup.control;
23726
2ebecff57b17 Basic editing of theory sources.
wenzelm
parents:
diff changeset
    60
37192
8cdddd689ea9 markup non-identifier keyword as operator;
wenzelm
parents: 36959
diff changeset
    61
fun token_markup tok =
45666
d83797ef0d2d separate module for concrete Isabelle markup;
wenzelm
parents: 44736
diff changeset
    62
  if Token.keyword_with (not o Lexicon.is_ascii_identifier) tok then Isabelle_Markup.operator
37197
953fc4983439 more detailed token markup, including command kind as sub_kind;
wenzelm
parents: 37193
diff changeset
    63
  else
953fc4983439 more detailed token markup, including command kind as sub_kind;
wenzelm
parents: 37193
diff changeset
    64
    let
953fc4983439 more detailed token markup, including command kind as sub_kind;
wenzelm
parents: 37193
diff changeset
    65
      val kind = Token.kind_of tok;
953fc4983439 more detailed token markup, including command kind as sub_kind;
wenzelm
parents: 37193
diff changeset
    66
      val props =
38471
0924654b8163 report command token name instead of kind, which can be retrieved later via Outer_Syntax.keyword_kind;
wenzelm
parents: 38422
diff changeset
    67
        if kind = Token.Command
0924654b8163 report command token name instead of kind, which can be retrieved later via Outer_Syntax.keyword_kind;
wenzelm
parents: 38422
diff changeset
    68
        then Markup.properties [(Markup.nameN, Token.content_of tok)]
37197
953fc4983439 more detailed token markup, including command kind as sub_kind;
wenzelm
parents: 37193
diff changeset
    69
        else I;
953fc4983439 more detailed token markup, including command kind as sub_kind;
wenzelm
parents: 37193
diff changeset
    70
    in props (token_kind_markup kind) end;
37192
8cdddd689ea9 markup non-identifier keyword as operator;
wenzelm
parents: 36959
diff changeset
    71
48749
c197b3c3e7fa some attempts to keep malformed syntax errors focussed, without too much red spilled onto the document view;
wenzelm
parents: 46811
diff changeset
    72
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
    73
  let
c197b3c3e7fa some attempts to keep malformed syntax errors focussed, without too much red spilled onto the document view;
wenzelm
parents: 46811
diff changeset
    74
    val malformed_symbols =
c197b3c3e7fa some attempts to keep malformed syntax errors focussed, without too much red spilled onto the document view;
wenzelm
parents: 46811
diff changeset
    75
      Symbol_Pos.explode (Token.source_position_of tok)
c197b3c3e7fa some attempts to keep malformed syntax errors focussed, without too much red spilled onto the document view;
wenzelm
parents: 46811
diff changeset
    76
      |> map_filter (fn (sym, pos) =>
48751
dc3bbdda4bc8 more visible markup of malformed input as "bad";
wenzelm
parents: 48749
diff changeset
    77
          if Symbol.is_malformed sym then SOME (pos, Isabelle_Markup.bad) 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
    78
    val is_malformed = Token.is_error tok orelse not (null malformed_symbols);
c197b3c3e7fa some attempts to keep malformed syntax errors focussed, without too much red spilled onto the document view;
wenzelm
parents: 46811
diff changeset
    79
    val reports = (Token.position_of tok, token_markup tok) :: malformed_symbols;
c197b3c3e7fa some attempts to keep malformed syntax errors focussed, without too much red spilled onto the document view;
wenzelm
parents: 46811
diff changeset
    80
  in (is_malformed, reports) end;
40528
d5e9f7608341 report malformed symbols;
wenzelm
parents: 40523
diff changeset
    81
23803
11bf7af10ec8 tuned signature;
wenzelm
parents: 23790
diff changeset
    82
in
11bf7af10ec8 tuned signature;
wenzelm
parents: 23790
diff changeset
    83
48749
c197b3c3e7fa some attempts to keep malformed syntax errors focussed, without too much red spilled onto the document view;
wenzelm
parents: 46811
diff changeset
    84
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
    85
  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
    86
  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
    87
23726
2ebecff57b17 Basic editing of theory sources.
wenzelm
parents:
diff changeset
    88
fun present_token tok =
37192
8cdddd689ea9 markup non-identifier keyword as operator;
wenzelm
parents: 36959
diff changeset
    89
  Markup.enclose (token_markup tok) (Output.output (Token.unparse tok));
23726
2ebecff57b17 Basic editing of theory sources.
wenzelm
parents:
diff changeset
    90
23803
11bf7af10ec8 tuned signature;
wenzelm
parents: 23790
diff changeset
    91
end;
11bf7af10ec8 tuned signature;
wenzelm
parents: 23790
diff changeset
    92
11bf7af10ec8 tuned signature;
wenzelm
parents: 23790
diff changeset
    93
11bf7af10ec8 tuned signature;
wenzelm
parents: 23790
diff changeset
    94
27665
b9e54ba563b3 renamed item to span, renamed contructors;
wenzelm
parents: 27353
diff changeset
    95
(** spans **)
b9e54ba563b3 renamed item to span, renamed contructors;
wenzelm
parents: 27353
diff changeset
    96
27842
6c35d50d309f abstract type span, tuned interfaces;
wenzelm
parents: 27770
diff changeset
    97
(* type span *)
6c35d50d309f abstract type span, tuned interfaces;
wenzelm
parents: 27770
diff changeset
    98
6c35d50d309f abstract type span, tuned interfaces;
wenzelm
parents: 27770
diff changeset
    99
datatype span_kind = Command of string | Ignored | Malformed;
36959
f5417836dbea renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
wenzelm
parents: 36950
diff changeset
   100
datatype span = Span of span_kind * Token.T list;
23803
11bf7af10ec8 tuned signature;
wenzelm
parents: 23790
diff changeset
   101
27842
6c35d50d309f abstract type span, tuned interfaces;
wenzelm
parents: 27770
diff changeset
   102
fun span_kind (Span (k, _)) = k;
6c35d50d309f abstract type span, tuned interfaces;
wenzelm
parents: 27770
diff changeset
   103
fun span_content (Span (_, toks)) = toks;
6c35d50d309f abstract type span, tuned interfaces;
wenzelm
parents: 27770
diff changeset
   104
23803
11bf7af10ec8 tuned signature;
wenzelm
parents: 23790
diff changeset
   105
11bf7af10ec8 tuned signature;
wenzelm
parents: 23790
diff changeset
   106
(* parse *)
23726
2ebecff57b17 Basic editing of theory sources.
wenzelm
parents:
diff changeset
   107
23803
11bf7af10ec8 tuned signature;
wenzelm
parents: 23790
diff changeset
   108
local
11bf7af10ec8 tuned signature;
wenzelm
parents: 23790
diff changeset
   109
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
   110
fun make_span toks =
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
   111
  if not (null toks) andalso Token.is_command (hd toks) then
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
   112
    Span (Command (Token.content_of (hd toks)), toks)
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
   113
  else if forall (not o Token.is_proper) toks then Span (Ignored, toks)
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
   114
  else Span (Malformed, toks);
23726
2ebecff57b17 Basic editing of theory sources.
wenzelm
parents:
diff changeset
   115
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
   116
fun flush (result, span) = if null span then (result, span) else (rev span :: result, []);
23726
2ebecff57b17 Basic editing of theory sources.
wenzelm
parents:
diff changeset
   117
2ebecff57b17 Basic editing of theory sources.
wenzelm
parents:
diff changeset
   118
in
2ebecff57b17 Basic editing of theory sources.
wenzelm
parents:
diff changeset
   119
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
   120
fun parse_spans toks =
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
   121
  fold (fn tok => Token.is_command tok ? flush #> apsnd (cons tok)) toks ([], [])
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
   122
  |> flush
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
   123
  |> #1 |> rev |> map make_span;
23803
11bf7af10ec8 tuned signature;
wenzelm
parents: 23790
diff changeset
   124
11bf7af10ec8 tuned signature;
wenzelm
parents: 23790
diff changeset
   125
end;
11bf7af10ec8 tuned signature;
wenzelm
parents: 23790
diff changeset
   126
11bf7af10ec8 tuned signature;
wenzelm
parents: 23790
diff changeset
   127
11bf7af10ec8 tuned signature;
wenzelm
parents: 23790
diff changeset
   128
(* present *)
11bf7af10ec8 tuned signature;
wenzelm
parents: 23790
diff changeset
   129
48752
8a81ef0bc790 discontinued mostly unused markup for command spans;
wenzelm
parents: 48751
diff changeset
   130
val present_span = implode o map present_token o span_content;
23803
11bf7af10ec8 tuned signature;
wenzelm
parents: 23790
diff changeset
   131
28434
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
   132
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
   133
43621
e8858190cccd clarified Thy_Syntax.element;
wenzelm
parents: 43430
diff changeset
   134
(** specification elements: commands with optional proof **)
e8858190cccd clarified Thy_Syntax.element;
wenzelm
parents: 43430
diff changeset
   135
e8858190cccd clarified Thy_Syntax.element;
wenzelm
parents: 43430
diff changeset
   136
type element = {head: span, proof: span list, proper_proof: bool};
e8858190cccd clarified Thy_Syntax.element;
wenzelm
parents: 43430
diff changeset
   137
e8858190cccd clarified Thy_Syntax.element;
wenzelm
parents: 43430
diff changeset
   138
fun make_element head proof proper_proof =
e8858190cccd clarified Thy_Syntax.element;
wenzelm
parents: 43430
diff changeset
   139
  {head = head, proof = proof, proper_proof = proper_proof};
e8858190cccd clarified Thy_Syntax.element;
wenzelm
parents: 43430
diff changeset
   140
28434
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
   141
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
   142
(* scanning spans *)
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
   143
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
   144
val eof = Span (Command "", []);
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
   145
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
   146
fun is_eof (Span (Command "", _)) = true
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
   147
  | is_eof _ = false;
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
   148
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
   149
val not_eof = not o is_eof;
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
   150
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
   151
val stopper = Scan.stopper (K eof) is_eof;
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
   152
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
   153
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
   154
(* parse *)
28434
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
   155
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
   156
local
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
   157
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
   158
fun command_with pred = Scan.one (fn (Span (Command name, _)) => pred name | _ => false);
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
   159
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
   160
val proof = Scan.pass 1 (Scan.repeat (Scan.depend (fn d =>
28454
c63168db774c unit_source: more rigid parsing, stop after final qed;
wenzelm
parents: 28438
diff changeset
   161
  if d <= 0 then Scan.fail
28434
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
   162
  else
36950
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 30573
diff changeset
   163
    command_with Keyword.is_qed_global >> pair ~1 ||
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 30573
diff changeset
   164
    command_with Keyword.is_proof_goal >> pair (d + 1) ||
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 30573
diff changeset
   165
    (if d = 0 then Scan.fail else command_with Keyword.is_qed >> pair (d - 1)) ||
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 30573
diff changeset
   166
    Scan.unless (command_with Keyword.is_theory) (Scan.one not_eof) >> pair d)) -- Scan.state);
28434
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
   167
43621
e8858190cccd clarified Thy_Syntax.element;
wenzelm
parents: 43430
diff changeset
   168
val element =
e8858190cccd clarified Thy_Syntax.element;
wenzelm
parents: 43430
diff changeset
   169
  command_with Keyword.is_theory_goal -- proof
e8858190cccd clarified Thy_Syntax.element;
wenzelm
parents: 43430
diff changeset
   170
    >> (fn (a, (bs, d)) => make_element a bs (d >= 0)) ||
e8858190cccd clarified Thy_Syntax.element;
wenzelm
parents: 43430
diff changeset
   171
  Scan.one not_eof >> (fn a => make_element a [] true);
28434
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
   172
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
   173
in
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
   174
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
   175
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
   176
  Source.of_list #>
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
   177
  Source.source stopper (Scan.bulk element) NONE #>
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
   178
  Source.exhaust;
28434
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
   179
23726
2ebecff57b17 Basic editing of theory sources.
wenzelm
parents:
diff changeset
   180
end;
28434
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
   181
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
   182
end;