src/Pure/ProofGeneral/pgip_parser.ML
author wenzelm
Sat, 13 Dec 2008 15:00:39 +0100
changeset 29091 b81fe045e799
parent 28035 7120e58464e4
child 29315 b074c05f00ad
permissions -rw-r--r--
Context.display_names;
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
    ID:         $Id$
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
     3
    Author:     David Aspinall and Makarius
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
     4
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
     5
Parsing theory sources without execution (via keyword classification).
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
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
     8
signature PGIP_PARSER =
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
     9
sig
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    10
  val pgip_parser: Position.T -> string -> PgipMarkup.pgipdocument
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    11
end
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    12
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    13
structure PgipParser: PGIP_PARSER =
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    14
struct
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    15
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    16
structure K = OuterKeyword;
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    17
structure D = PgipMarkup;
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    18
structure I = PgipIsabelle;
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    19
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 badcmd text = [D.Badcmd {text = text}];
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    22
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    23
fun thy_begin text =
23868
8c6f2e7bfdb4 adapted ThyHeader.read;
wenzelm
parents: 23797
diff changeset
    24
  (case try (ThyHeader.read Position.none) (Source.of_string text) of
24192
4eccd4bb8b64 PGIP change: thyname is optional in opentheory, markup even in case of header parse failure
aspinall
parents: 23868
diff changeset
    25
    NONE => D.Opentheory {thyname = NONE, parentnames = [], text = text}
23797
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    26
  | 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
    27
       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
    28
  :: [D.Openblock {metavarid = NONE, name = NONE, objtype = SOME I.ObjTheoryBody}];
23797
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    29
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    30
fun thy_heading text =
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    31
  [D.Closeblock {},
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    32
   D.Doccomment {text = text},
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    33
   D.Openblock {metavarid = NONE, name = NONE, objtype = SOME I.ObjTheoryBody}];
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    34
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    35
fun thy_decl text =
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    36
  [D.Theoryitem {name = NONE, objtype = SOME I.ObjTheoryDecl, text = text}];
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    37
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    38
fun goal text =
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    39
  [D.Opengoal {thmname = NONE, text = text},
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    40
   D.Openblock {metavarid = NONE, name = NONE, objtype = SOME I.ObjProofBody}];
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    41
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    42
fun prf_block text =
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    43
  [D.Closeblock {},
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    44
   D.Proofstep {text = text},
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    45
   D.Openblock {metavarid = NONE, name = NONE, objtype = SOME I.ObjProofBody}];
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    46
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    47
fun prf_open text =
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    48
 [D.Openblock {metavarid = NONE, name = NONE, objtype = SOME I.ObjProofBody},
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    49
  D.Proofstep {text = text}];
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    50
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    51
fun proofstep text = [D.Proofstep {text = text}];
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    52
fun closegoal text = [D.Closegoal {text = text}, D.Closeblock {}];
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    53
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
fun command k f = Symtab.update_new (OuterKeyword.kind_of k, f);
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    56
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    57
val command_keywords = Symtab.empty
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    58
  |> command K.control          badcmd
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    59
  |> command K.diag             (fn text => [D.Spuriouscmd {text = text}])
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    60
  |> command K.thy_begin        thy_begin
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    61
  |> command K.thy_switch       badcmd
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    62
  |> command K.thy_end          (fn text => [D.Closeblock {}, D.Closetheory {text = text}])
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    63
  |> command K.thy_heading      thy_heading
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    64
  |> command K.thy_decl         thy_decl
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    65
  |> command K.thy_script       thy_decl
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    66
  |> command K.thy_goal         goal
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    67
  |> command K.qed              closegoal
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    68
  |> command K.qed_block        closegoal
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    69
  |> command K.qed_global       (fn text => [D.Giveupgoal {text = text}])
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    70
  |> command K.prf_heading      (fn text => [D.Doccomment {text = text}])
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    71
  |> command K.prf_goal         goal
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    72
  |> command K.prf_block        prf_block
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    73
  |> command K.prf_open         prf_open
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    74
  |> command K.prf_close        (fn text => [D.Proofstep {text = text}, D.Closeblock {}])
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    75
  |> command K.prf_chain        proofstep
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    76
  |> command K.prf_decl         proofstep
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    77
  |> command K.prf_asm          proofstep
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    78
  |> command K.prf_asm_goal     goal
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    79
  |> command K.prf_script       proofstep;
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    80
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    81
val _ = OuterKeyword.kinds subset_string Symtab.keys command_keywords
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    82
  orelse sys_error "Incomplete coverage of command keywords";
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    83
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    84
fun parse_command "sorry" text = [D.Postponegoal {text = text}, D.Closeblock {}]
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    85
  | parse_command name text =
27353
71c4dd53d4cb moved global keywords from OuterSyntax to OuterKeyword, tuned interfaces;
wenzelm
parents: 24192
diff changeset
    86
      (case OuterKeyword.command_keyword name of
23797
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    87
        NONE => [D.Unparseable {text = text}]
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    88
      | SOME k =>
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    89
          (case Symtab.lookup command_keywords (OuterKeyword.kind_of k) of
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 f => f text));
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    92
27841
55b028593335 adapted ThyEdit operations;
wenzelm
parents: 27664
diff changeset
    93
fun parse_span span =
55b028593335 adapted ThyEdit operations;
wenzelm
parents: 27664
diff changeset
    94
  let
55b028593335 adapted ThyEdit operations;
wenzelm
parents: 27664
diff changeset
    95
    val kind = ThyEdit.span_kind span;
55b028593335 adapted ThyEdit operations;
wenzelm
parents: 27664
diff changeset
    96
    val toks = ThyEdit.span_content span;
28035
7120e58464e4 present_token: disable print_mode, which is YXML now;
wenzelm
parents: 27841
diff changeset
    97
    val text = implode (map (PrintMode.setmp [] ThyEdit.present_token) toks);
27841
55b028593335 adapted ThyEdit operations;
wenzelm
parents: 27664
diff changeset
    98
  in
23797
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
    99
    (case kind of
27664
0e7a7fcd403b adapted ThyEdit.span;
wenzelm
parents: 27353
diff changeset
   100
      ThyEdit.Command name => parse_command name text
0e7a7fcd403b adapted ThyEdit.span;
wenzelm
parents: 27353
diff changeset
   101
    | ThyEdit.Ignored => [D.Whitespace {text = text}]
27841
55b028593335 adapted ThyEdit operations;
wenzelm
parents: 27664
diff changeset
   102
    | ThyEdit.Malformed => [D.Unparseable {text = text}])
23797
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
   103
  end;
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
   104
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
   105
f4dbbffbfe06 Parsing theory sources without execution (via keyword classification).
wenzelm
parents:
diff changeset
   106
fun pgip_parser pos str =
27841
55b028593335 adapted ThyEdit operations;
wenzelm
parents: 27664
diff changeset
   107
  maps parse_span (ThyEdit.parse_spans (OuterKeyword.get_lexicons ()) pos str);
23797
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
end;