src/Pure/ProofGeneral/pgip_parser.ML
author wenzelm
Mon, 19 Mar 2012 18:18:42 +0100
changeset 47014 e203b7d7e08d
parent 46969 481b7d9ad6fe
child 48864 3ee314ae1e0a
permissions -rw-r--r--
allow keyword tags to be redefined, but not the command category;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
23797
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/ProofGeneral/pgip_parser.ML
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
     2
    Author:     David Aspinall and Makarius
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
     3
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
     4
Parsing theory sources without execution (via keyword classification).
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
     5
*)
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
     6
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
     7
signature PGIP_PARSER =
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
     8
sig
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
     9
  val pgip_parser: Position.T -> string -> PgipMarkup.pgipdocument
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    10
end
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    11
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    12
structure PgipParser: PGIP_PARSER =
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    13
struct
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    14
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    15
structure D = PgipMarkup;
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    16
structure I = PgipIsabelle;
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    17
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    18
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    19
fun badcmd text = [D.Badcmd {text = text}];
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    20
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    21
fun thy_begin text =
37978
548f3f165d05 simplified Thy_Header.read -- include Source.of_string_limited here;
wenzelm
parents: 37852
diff changeset
    22
  (case try (Thy_Header.read Position.none) text of
24192
4eccd4bb8b64 PGIP change: thyname is optional in opentheory, markup even in case of header parse failure
aspinall
parents: 23868
diff changeset
    23
    NONE => D.Opentheory {thyname = NONE, parentnames = [], text = text}
46938
cda018294515 some support for outer syntax keyword declarations within theory header;
wenzelm
parents: 46811
diff changeset
    24
  | SOME {name, imports, ...} =>
24192
4eccd4bb8b64 PGIP change: thyname is optional in opentheory, markup even in case of header parse failure
aspinall
parents: 23868
diff changeset
    25
       D.Opentheory {thyname = SOME name, parentnames = imports, text = text})
4eccd4bb8b64 PGIP change: thyname is optional in opentheory, markup even in case of header parse failure
aspinall
parents: 23868
diff changeset
    26
  :: [D.Openblock {metavarid = NONE, name = NONE, objtype = SOME I.ObjTheoryBody}];
23797
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    27
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    28
fun thy_heading text =
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    29
  [D.Closeblock {},
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    30
   D.Doccomment {text = text},
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    31
   D.Openblock {metavarid = NONE, name = NONE, objtype = SOME I.ObjTheoryBody}];
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    32
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    33
fun thy_decl text =
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    34
  [D.Theoryitem {name = NONE, objtype = SOME I.ObjTheoryDecl, text = text}];
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    35
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    36
fun goal text =
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    37
  [D.Opengoal {thmname = NONE, text = text},
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    38
   D.Openblock {metavarid = NONE, name = NONE, objtype = SOME I.ObjProofBody}];
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    39
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    40
fun prf_block text =
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    41
  [D.Closeblock {},
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    42
   D.Proofstep {text = text},
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    43
   D.Openblock {metavarid = NONE, name = NONE, objtype = SOME I.ObjProofBody}];
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    44
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    45
fun prf_open text =
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    46
 [D.Openblock {metavarid = NONE, name = NONE, objtype = SOME I.ObjProofBody},
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    47
  D.Proofstep {text = text}];
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    48
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    49
fun proofstep text = [D.Proofstep {text = text}];
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    50
fun closegoal text = [D.Closegoal {text = text}, D.Closeblock {}];
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    51
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    52
36950
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 36315
diff changeset
    53
fun command k f = Symtab.update_new (Keyword.kind_of k, f);
23797
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    54
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    55
val command_keywords = Symtab.empty
36950
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 36315
diff changeset
    56
  |> command Keyword.control          badcmd
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 36315
diff changeset
    57
  |> command Keyword.diag             (fn text => [D.Spuriouscmd {text = text}])
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 36315
diff changeset
    58
  |> command Keyword.thy_begin        thy_begin
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 36315
diff changeset
    59
  |> command Keyword.thy_end          (fn text => [D.Closeblock {}, D.Closetheory {text = text}])
46969
481b7d9ad6fe more abstract heading level;
wenzelm
parents: 46968
diff changeset
    60
  |> command Keyword.thy_heading1     thy_heading
481b7d9ad6fe more abstract heading level;
wenzelm
parents: 46968
diff changeset
    61
  |> command Keyword.thy_heading2     thy_heading
481b7d9ad6fe more abstract heading level;
wenzelm
parents: 46968
diff changeset
    62
  |> command Keyword.thy_heading3     thy_heading
481b7d9ad6fe more abstract heading level;
wenzelm
parents: 46968
diff changeset
    63
  |> command Keyword.thy_heading4     thy_heading
36950
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 36315
diff changeset
    64
  |> command Keyword.thy_decl         thy_decl
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 36315
diff changeset
    65
  |> command Keyword.thy_script       thy_decl
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 36315
diff changeset
    66
  |> command Keyword.thy_goal         goal
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 36315
diff changeset
    67
  |> command Keyword.thy_schematic_goal goal
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 36315
diff changeset
    68
  |> command Keyword.qed              closegoal
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 36315
diff changeset
    69
  |> command Keyword.qed_block        closegoal
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 36315
diff changeset
    70
  |> command Keyword.qed_global       (fn text => [D.Giveupgoal {text = text}])
46969
481b7d9ad6fe more abstract heading level;
wenzelm
parents: 46968
diff changeset
    71
  |> command Keyword.prf_heading2     (fn text => [D.Doccomment {text = text}])
481b7d9ad6fe more abstract heading level;
wenzelm
parents: 46968
diff changeset
    72
  |> command Keyword.prf_heading3     (fn text => [D.Doccomment {text = text}])
481b7d9ad6fe more abstract heading level;
wenzelm
parents: 46968
diff changeset
    73
  |> command Keyword.prf_heading4     (fn text => [D.Doccomment {text = text}])
36950
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 36315
diff changeset
    74
  |> command Keyword.prf_goal         goal
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 36315
diff changeset
    75
  |> command Keyword.prf_block        prf_block
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 36315
diff changeset
    76
  |> command Keyword.prf_open         prf_open
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 36315
diff changeset
    77
  |> command Keyword.prf_close        (fn text => [D.Proofstep {text = text}, D.Closeblock {}])
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 36315
diff changeset
    78
  |> command Keyword.prf_chain        proofstep
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 36315
diff changeset
    79
  |> command Keyword.prf_decl         proofstep
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 36315
diff changeset
    80
  |> command Keyword.prf_asm          proofstep
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 36315
diff changeset
    81
  |> command Keyword.prf_asm_goal     goal
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 36315
diff changeset
    82
  |> command Keyword.prf_script       proofstep;
23797
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    83
46968
38aaa08fb37f less redundant data;
wenzelm
parents: 46967
diff changeset
    84
val _ = subset (op =) (map Keyword.kind_of Keyword.kinds, Symtab.keys command_keywords)
37852
a902f158b4fc eliminated old-style sys_error/SYS_ERROR in favour of exception Fail -- after careful checking that there is no overlap with existing handling of that;
wenzelm
parents: 37216
diff changeset
    85
  orelse raise Fail "Incomplete coverage of command keywords";
23797
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    86
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    87
fun parse_command "sorry" text = [D.Postponegoal {text = text}, D.Closeblock {}]
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    88
  | parse_command name text =
36950
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 36315
diff changeset
    89
      (case Keyword.command_keyword name of
23797
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    90
        NONE => [D.Unparseable {text = text}]
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    91
      | SOME k =>
36950
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 36315
diff changeset
    92
          (case Symtab.lookup command_keywords (Keyword.kind_of k) of
23797
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    93
            NONE => [D.Unparseable {text = text}]
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    94
          | SOME f => f text));
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    95
27841
55b028593335 adapted ThyEdit operations;
wenzelm
parents: 27664
diff changeset
    96
fun parse_span span =
55b028593335 adapted ThyEdit operations;
wenzelm
parents: 27664
diff changeset
    97
  let
37216
3165bc303f66 modernized some structure names, keeping a few legacy aliases;
wenzelm
parents: 37146
diff changeset
    98
    val kind = Thy_Syntax.span_kind span;
3165bc303f66 modernized some structure names, keeping a few legacy aliases;
wenzelm
parents: 37146
diff changeset
    99
    val toks = Thy_Syntax.span_content span;
3165bc303f66 modernized some structure names, keeping a few legacy aliases;
wenzelm
parents: 37146
diff changeset
   100
    val text = implode (map (Print_Mode.setmp [] Thy_Syntax.present_token) toks);
27841
55b028593335 adapted ThyEdit operations;
wenzelm
parents: 27664
diff changeset
   101
  in
23797
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
   102
    (case kind of
37216
3165bc303f66 modernized some structure names, keeping a few legacy aliases;
wenzelm
parents: 37146
diff changeset
   103
      Thy_Syntax.Command name => parse_command name text
3165bc303f66 modernized some structure names, keeping a few legacy aliases;
wenzelm
parents: 37146
diff changeset
   104
    | Thy_Syntax.Ignored => [D.Whitespace {text = text}]
3165bc303f66 modernized some structure names, keeping a few legacy aliases;
wenzelm
parents: 37146
diff changeset
   105
    | Thy_Syntax.Malformed => [D.Unparseable {text = text}])
23797
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
   106
  end;
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
   107
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
   108
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
   109
fun pgip_parser pos str =
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: 37978
diff changeset
   110
  Thy_Syntax.parse_tokens (Keyword.get_lexicons ()) pos str
03a2dc9e0624 clarified command span: include trailing whitespace/comments and thus reduce number of ignored spans with associated transactions and states (factor 2);
wenzelm
parents: 37978
diff changeset
   111
  |> Thy_Syntax.parse_spans
03a2dc9e0624 clarified command span: include trailing whitespace/comments and thus reduce number of ignored spans with associated transactions and states (factor 2);
wenzelm
parents: 37978
diff changeset
   112
  |> maps parse_span;
23797
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
   113
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
   114
end;