src/Pure/ML/ml_parse.ML
author boehmes
Tue, 07 Dec 2010 14:54:31 +0100
changeset 41061 492f8fd35fc0
parent 30682 dcb233670c98
child 43947 9b00f09f7721
permissions -rw-r--r--
centralized handling of built-in types and constants for bitvectors
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
24594
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/ML/ml_parse.ML
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
     2
    Author:     Makarius
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
     3
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
     4
Minimal parsing for SML -- fixing integer numerals.
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
     5
*)
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
     6
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
     7
signature ML_PARSE =
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
     8
sig
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
     9
  val fix_ints: string -> string
30672
beaadd5af500 more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
wenzelm
parents: 30636
diff changeset
    10
  val global_context: use_context
24594
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
    11
end;
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
    12
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
    13
structure ML_Parse: ML_PARSE =
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
    14
struct
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
    15
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
    16
(** error handling **)
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
    17
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
    18
fun !!! scan =
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
    19
  let
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
    20
    fun get_pos [] = " (past end-of-file!)"
30682
dcb233670c98 eliminated non-canonical alias structure T = ML_Lex;
wenzelm
parents: 30672
diff changeset
    21
      | get_pos (tok :: _) = Position.str_of (ML_Lex.pos_of tok);
24594
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
    22
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
    23
    fun err (toks, NONE) = "SML syntax error" ^ get_pos toks
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
    24
      | err (toks, SOME msg) = "SML syntax error" ^ get_pos toks ^ ": " ^ msg;
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
    25
  in Scan.!! err scan end;
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
    26
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
    27
fun bad_input x =
30682
dcb233670c98 eliminated non-canonical alias structure T = ML_Lex;
wenzelm
parents: 30672
diff changeset
    28
  (Scan.some (fn tok => (case ML_Lex.kind_of tok of ML_Lex.Error msg => SOME msg | _ => NONE)) :|--
24594
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
    29
    (fn msg => Scan.fail_with (K msg))) x;
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
    30
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
    31
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
    32
(** basic parsers **)
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
    33
27817
78cae5cca09e renamed ML_Lex.val_of to content_of;
wenzelm
parents: 24594
diff changeset
    34
fun $$$ x =
30682
dcb233670c98 eliminated non-canonical alias structure T = ML_Lex;
wenzelm
parents: 30672
diff changeset
    35
  Scan.one (fn tok => ML_Lex.kind_of tok = ML_Lex.Keyword andalso ML_Lex.content_of tok = x)
dcb233670c98 eliminated non-canonical alias structure T = ML_Lex;
wenzelm
parents: 30672
diff changeset
    36
    >> ML_Lex.content_of;
24594
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
    37
30682
dcb233670c98 eliminated non-canonical alias structure T = ML_Lex;
wenzelm
parents: 30672
diff changeset
    38
val int = Scan.one (fn tok => ML_Lex.kind_of tok = ML_Lex.Int) >> ML_Lex.content_of;
dcb233670c98 eliminated non-canonical alias structure T = ML_Lex;
wenzelm
parents: 30672
diff changeset
    39
dcb233670c98 eliminated non-canonical alias structure T = ML_Lex;
wenzelm
parents: 30672
diff changeset
    40
val regular = Scan.one ML_Lex.is_regular >> ML_Lex.content_of;
dcb233670c98 eliminated non-canonical alias structure T = ML_Lex;
wenzelm
parents: 30672
diff changeset
    41
val improper = Scan.one ML_Lex.is_improper >> ML_Lex.content_of;
24594
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
    42
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
    43
val blanks = Scan.repeat improper >> implode;
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
    44
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
    45
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
    46
(* fix_ints *)
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
    47
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
    48
(*approximation only -- corrupts numeric record field patterns *)
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
    49
val fix_int =
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
    50
  $$$ "#" ^^ blanks ^^ int ||
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
    51
  ($$$ "infix" || $$$ "infixr") ^^ blanks ^^ int ||
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
    52
  int >> (fn x => "(" ^ x ^ ":int)") ||
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
    53
  regular ||
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
    54
  bad_input;
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
    55
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
    56
fun do_fix_ints s =
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
    57
  Source.of_string s
30682
dcb233670c98 eliminated non-canonical alias structure T = ML_Lex;
wenzelm
parents: 30672
diff changeset
    58
  |> ML_Lex.source
dcb233670c98 eliminated non-canonical alias structure T = ML_Lex;
wenzelm
parents: 30672
diff changeset
    59
  |> Source.source ML_Lex.stopper (Scan.bulk (!!! fix_int)) NONE
24594
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
    60
  |> Source.exhaust
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
    61
  |> implode;
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
    62
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
    63
val fix_ints = if ml_system_fix_ints then do_fix_ints else I;
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
    64
30672
beaadd5af500 more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
wenzelm
parents: 30636
diff changeset
    65
beaadd5af500 more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
wenzelm
parents: 30636
diff changeset
    66
(* global use_context *)
beaadd5af500 more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
wenzelm
parents: 30636
diff changeset
    67
beaadd5af500 more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
wenzelm
parents: 30636
diff changeset
    68
val global_context: use_context =
beaadd5af500 more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
wenzelm
parents: 30636
diff changeset
    69
 {tune_source = fix_ints,
beaadd5af500 more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
wenzelm
parents: 30636
diff changeset
    70
  name_space = ML_Name_Space.global,
beaadd5af500 more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
wenzelm
parents: 30636
diff changeset
    71
  str_of_pos = Position.str_of oo Position.line_file,
beaadd5af500 more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
wenzelm
parents: 30636
diff changeset
    72
  print = writeln,
beaadd5af500 more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
wenzelm
parents: 30636
diff changeset
    73
  error = error};
beaadd5af500 more systematic type use_context, with particular values ML_Parse.global_context and ML_Context.local_context;
wenzelm
parents: 30636
diff changeset
    74
24594
6689effe75d1 Minimal parsing for SML -- fixing integer numerals.
wenzelm
parents:
diff changeset
    75
end;