src/Pure/Thy/thy_syntax.ML
author wenzelm
Fri, 09 Oct 2015 21:16:00 +0200
changeset 61379 c57820ceead3
parent 59809 87641097d0f3
permissions -rw-r--r--
more direct HTML presentation, without print mode;
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
61379
c57820ceead3 more direct HTML presentation, without print mode;
wenzelm
parents: 59809
diff changeset
     4
Theory syntax elements.
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
51225
3fe0d8d55975 support nested Thy_Syntax.element;
wenzelm
parents: 50238
diff changeset
     9
  datatype 'a element = Element of 'a * ('a element list * 'a) option
51321
75682dfff630 more Thy_Syntax.element operations;
wenzelm
parents: 51268
diff changeset
    10
  val atom: 'a -> 'a element
51225
3fe0d8d55975 support nested Thy_Syntax.element;
wenzelm
parents: 50238
diff changeset
    11
  val map_element: ('a -> 'b) -> 'a element -> 'b element
3fe0d8d55975 support nested Thy_Syntax.element;
wenzelm
parents: 50238
diff changeset
    12
  val flat_element: 'a element -> 'a list
51268
fcc4b89a600d simplified Outer_Syntax.read_span: internalized Toplevel.is_ignored;
wenzelm
parents: 51267
diff changeset
    13
  val last_element: 'a element -> 'a
58923
cb9b69cca999 more explicit Keyword.keywords;
wenzelm
parents: 58864
diff changeset
    14
  val parse_elements: Keyword.keywords -> Command_Span.span list -> Command_Span.span element list
23726
2ebecff57b17 Basic editing of theory sources.
wenzelm
parents:
diff changeset
    15
end;
2ebecff57b17 Basic editing of theory sources.
wenzelm
parents:
diff changeset
    16
37216
3165bc303f66 modernized some structure names, keeping a few legacy aliases;
wenzelm
parents: 37197
diff changeset
    17
structure Thy_Syntax: THY_SYNTAX =
23726
2ebecff57b17 Basic editing of theory sources.
wenzelm
parents:
diff changeset
    18
struct
2ebecff57b17 Basic editing of theory sources.
wenzelm
parents:
diff changeset
    19
61379
c57820ceead3 more direct HTML presentation, without print mode;
wenzelm
parents: 59809
diff changeset
    20
(* datatype element: command with optional proof *)
43621
e8858190cccd clarified Thy_Syntax.element;
wenzelm
parents: 43430
diff changeset
    21
51225
3fe0d8d55975 support nested Thy_Syntax.element;
wenzelm
parents: 50238
diff changeset
    22
datatype 'a element = Element of 'a * ('a element list * 'a) option;
3fe0d8d55975 support nested Thy_Syntax.element;
wenzelm
parents: 50238
diff changeset
    23
3fe0d8d55975 support nested Thy_Syntax.element;
wenzelm
parents: 50238
diff changeset
    24
fun element (a, b) = Element (a, SOME b);
3fe0d8d55975 support nested Thy_Syntax.element;
wenzelm
parents: 50238
diff changeset
    25
fun atom a = Element (a, NONE);
43621
e8858190cccd clarified Thy_Syntax.element;
wenzelm
parents: 43430
diff changeset
    26
51225
3fe0d8d55975 support nested Thy_Syntax.element;
wenzelm
parents: 50238
diff changeset
    27
fun map_element f (Element (a, NONE)) = Element (f a, NONE)
3fe0d8d55975 support nested Thy_Syntax.element;
wenzelm
parents: 50238
diff changeset
    28
  | map_element f (Element (a, SOME (elems, b))) =
3fe0d8d55975 support nested Thy_Syntax.element;
wenzelm
parents: 50238
diff changeset
    29
      Element (f a, SOME ((map o map_element) f elems, f b));
3fe0d8d55975 support nested Thy_Syntax.element;
wenzelm
parents: 50238
diff changeset
    30
3fe0d8d55975 support nested Thy_Syntax.element;
wenzelm
parents: 50238
diff changeset
    31
fun flat_element (Element (a, NONE)) = [a]
3fe0d8d55975 support nested Thy_Syntax.element;
wenzelm
parents: 50238
diff changeset
    32
  | flat_element (Element (a, SOME (elems, b))) = a :: maps flat_element elems @ [b];
43621
e8858190cccd clarified Thy_Syntax.element;
wenzelm
parents: 43430
diff changeset
    33
51268
fcc4b89a600d simplified Outer_Syntax.read_span: internalized Toplevel.is_ignored;
wenzelm
parents: 51267
diff changeset
    34
fun last_element (Element (a, NONE)) = a
fcc4b89a600d simplified Outer_Syntax.read_span: internalized Toplevel.is_ignored;
wenzelm
parents: 51267
diff changeset
    35
  | last_element (Element (_, SOME (_, b))) = b;
fcc4b89a600d simplified Outer_Syntax.read_span: internalized Toplevel.is_ignored;
wenzelm
parents: 51267
diff changeset
    36
28434
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
    37
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
    38
(* scanning spans *)
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
    39
57905
c0c5652e796e separate module Command_Span: mostly syntactic representation;
wenzelm
parents: 57903
diff changeset
    40
val eof = Command_Span.Span (Command_Span.Command_Span ("", Position.none), []);
28434
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
    41
57905
c0c5652e796e separate module Command_Span: mostly syntactic representation;
wenzelm
parents: 57903
diff changeset
    42
fun is_eof (Command_Span.Span (Command_Span.Command_Span ("", _), _)) = true
28434
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
    43
  | is_eof _ = false;
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
    44
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
    45
val not_eof = not o is_eof;
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
    46
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
    47
val stopper = Scan.stopper (K eof) is_eof;
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
    48
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
    49
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
    50
(* parse *)
28434
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
    51
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
    52
local
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
    53
48878
5e850e6fa3c3 tuned errors;
wenzelm
parents: 48867
diff changeset
    54
fun command_with pred =
57905
c0c5652e796e separate module Command_Span: mostly syntactic representation;
wenzelm
parents: 57903
diff changeset
    55
  Scan.one
c0c5652e796e separate module Command_Span: mostly syntactic representation;
wenzelm
parents: 57903
diff changeset
    56
    (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
    57
58923
cb9b69cca999 more explicit Keyword.keywords;
wenzelm
parents: 58864
diff changeset
    58
fun parse_element keywords =
cb9b69cca999 more explicit Keyword.keywords;
wenzelm
parents: 58864
diff changeset
    59
  let
cb9b69cca999 more explicit Keyword.keywords;
wenzelm
parents: 58864
diff changeset
    60
    val proof_atom =
cb9b69cca999 more explicit Keyword.keywords;
wenzelm
parents: 58864
diff changeset
    61
      Scan.one
cb9b69cca999 more explicit Keyword.keywords;
wenzelm
parents: 58864
diff changeset
    62
        (fn (Command_Span.Span (Command_Span.Command_Span (name, _), _)) =>
cb9b69cca999 more explicit Keyword.keywords;
wenzelm
parents: 58864
diff changeset
    63
              Keyword.is_proof_body keywords name
cb9b69cca999 more explicit Keyword.keywords;
wenzelm
parents: 58864
diff changeset
    64
          | _ => true) >> atom;
cb9b69cca999 more explicit Keyword.keywords;
wenzelm
parents: 58864
diff changeset
    65
    fun proof_element x =
cb9b69cca999 more explicit Keyword.keywords;
wenzelm
parents: 58864
diff changeset
    66
      (command_with (Keyword.is_proof_goal keywords) -- proof_rest >> element || proof_atom) x
cb9b69cca999 more explicit Keyword.keywords;
wenzelm
parents: 58864
diff changeset
    67
    and proof_rest x =
cb9b69cca999 more explicit Keyword.keywords;
wenzelm
parents: 58864
diff changeset
    68
      (Scan.repeat proof_element -- command_with (Keyword.is_qed keywords)) x;
cb9b69cca999 more explicit Keyword.keywords;
wenzelm
parents: 58864
diff changeset
    69
  in
cb9b69cca999 more explicit Keyword.keywords;
wenzelm
parents: 58864
diff changeset
    70
    command_with (Keyword.is_theory_goal keywords) -- proof_rest >> element ||
cb9b69cca999 more explicit Keyword.keywords;
wenzelm
parents: 58864
diff changeset
    71
    Scan.one not_eof >> atom
cb9b69cca999 more explicit Keyword.keywords;
wenzelm
parents: 58864
diff changeset
    72
  end;
28434
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
    73
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
    74
in
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
    75
58923
cb9b69cca999 more explicit Keyword.keywords;
wenzelm
parents: 58864
diff changeset
    76
fun parse_elements keywords =
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
    77
  Source.of_list #>
58923
cb9b69cca999 more explicit Keyword.keywords;
wenzelm
parents: 58864
diff changeset
    78
  Source.source stopper (Scan.bulk (parse_element keywords)) #>
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
    79
  Source.exhaust;
28434
56f0951f4d26 added unit_source: commands with proof;
wenzelm
parents: 27863
diff changeset
    80
23726
2ebecff57b17 Basic editing of theory sources.
wenzelm
parents:
diff changeset
    81
end;
28434
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
end;