src/Pure/Tools/ghc.ML
author wenzelm
Tue, 30 Oct 2018 22:05:30 +0100
changeset 69211 7062639cfdaa
parent 69209 3f4210c13356
child 69216 1a52baa70aed
permissions -rw-r--r--
added GHC.read_source: read Haskell source text with antiquotations; added "cartouche" antiquotation for ML string expressions as Haskell string literals;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
69209
3f4210c13356 support for GHC: string literals;
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/Tools/ghc.ML
3f4210c13356 support for GHC: string literals;
wenzelm
parents:
diff changeset
     2
    Author:     Makarius
3f4210c13356 support for GHC: string literals;
wenzelm
parents:
diff changeset
     3
3f4210c13356 support for GHC: string literals;
wenzelm
parents:
diff changeset
     4
Support for GHC: Glasgow Haskell Compiler.
3f4210c13356 support for GHC: string literals;
wenzelm
parents:
diff changeset
     5
*)
3f4210c13356 support for GHC: string literals;
wenzelm
parents:
diff changeset
     6
3f4210c13356 support for GHC: string literals;
wenzelm
parents:
diff changeset
     7
signature GHC =
3f4210c13356 support for GHC: string literals;
wenzelm
parents:
diff changeset
     8
sig
3f4210c13356 support for GHC: string literals;
wenzelm
parents:
diff changeset
     9
  val print_codepoint: UTF8.codepoint -> string
3f4210c13356 support for GHC: string literals;
wenzelm
parents:
diff changeset
    10
  val print_symbol: Symbol.symbol -> string
3f4210c13356 support for GHC: string literals;
wenzelm
parents:
diff changeset
    11
  val print_string: string -> string
69211
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    12
  val antiquotation: binding -> 'a Token.context_parser ->
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    13
    ({context: Proof.context, source: Token.src, argument: 'a} -> string) -> theory -> theory
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    14
  val read_source: Proof.context -> Input.source -> string
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    15
  val set_result: string -> Proof.context -> Proof.context
69209
3f4210c13356 support for GHC: string literals;
wenzelm
parents:
diff changeset
    16
end;
3f4210c13356 support for GHC: string literals;
wenzelm
parents:
diff changeset
    17
3f4210c13356 support for GHC: string literals;
wenzelm
parents:
diff changeset
    18
structure GHC: GHC =
3f4210c13356 support for GHC: string literals;
wenzelm
parents:
diff changeset
    19
struct
3f4210c13356 support for GHC: string literals;
wenzelm
parents:
diff changeset
    20
3f4210c13356 support for GHC: string literals;
wenzelm
parents:
diff changeset
    21
(** string literals **)
3f4210c13356 support for GHC: string literals;
wenzelm
parents:
diff changeset
    22
3f4210c13356 support for GHC: string literals;
wenzelm
parents:
diff changeset
    23
fun print_codepoint c =
3f4210c13356 support for GHC: string literals;
wenzelm
parents:
diff changeset
    24
  (case c of
3f4210c13356 support for GHC: string literals;
wenzelm
parents:
diff changeset
    25
    34 => "\\\""
3f4210c13356 support for GHC: string literals;
wenzelm
parents:
diff changeset
    26
  | 39 => "\\'"
3f4210c13356 support for GHC: string literals;
wenzelm
parents:
diff changeset
    27
  | 92 => "\\\\"
3f4210c13356 support for GHC: string literals;
wenzelm
parents:
diff changeset
    28
  | 7 => "\\a"
3f4210c13356 support for GHC: string literals;
wenzelm
parents:
diff changeset
    29
  | 8 => "\\b"
3f4210c13356 support for GHC: string literals;
wenzelm
parents:
diff changeset
    30
  | 9 => "\\t"
3f4210c13356 support for GHC: string literals;
wenzelm
parents:
diff changeset
    31
  | 10 => "\\n"
3f4210c13356 support for GHC: string literals;
wenzelm
parents:
diff changeset
    32
  | 11 => "\\v"
3f4210c13356 support for GHC: string literals;
wenzelm
parents:
diff changeset
    33
  | 12 => "\\f"
3f4210c13356 support for GHC: string literals;
wenzelm
parents:
diff changeset
    34
  | 13 => "\\r"
3f4210c13356 support for GHC: string literals;
wenzelm
parents:
diff changeset
    35
  | c =>
3f4210c13356 support for GHC: string literals;
wenzelm
parents:
diff changeset
    36
      if c >= 32 andalso c < 127 then chr c
3f4210c13356 support for GHC: string literals;
wenzelm
parents:
diff changeset
    37
      else "\\" ^ string_of_int c ^ "\\&");
3f4210c13356 support for GHC: string literals;
wenzelm
parents:
diff changeset
    38
3f4210c13356 support for GHC: string literals;
wenzelm
parents:
diff changeset
    39
fun print_symbol sym =
3f4210c13356 support for GHC: string literals;
wenzelm
parents:
diff changeset
    40
  (case Symbol.decode sym of
3f4210c13356 support for GHC: string literals;
wenzelm
parents:
diff changeset
    41
    Symbol.Char s => print_codepoint (ord s)
3f4210c13356 support for GHC: string literals;
wenzelm
parents:
diff changeset
    42
  | Symbol.UTF8 s => UTF8.decode_permissive s |> map print_codepoint |> implode
3f4210c13356 support for GHC: string literals;
wenzelm
parents:
diff changeset
    43
  | Symbol.Sym s => "\\092<" ^ s ^ ">"
3f4210c13356 support for GHC: string literals;
wenzelm
parents:
diff changeset
    44
  | Symbol.Control s => "\\092<^" ^ s ^ ">"
3f4210c13356 support for GHC: string literals;
wenzelm
parents:
diff changeset
    45
  | _ => translate_string (print_codepoint o ord) sym);
3f4210c13356 support for GHC: string literals;
wenzelm
parents:
diff changeset
    46
3f4210c13356 support for GHC: string literals;
wenzelm
parents:
diff changeset
    47
val print_string = quote o implode o map print_symbol o Symbol.explode;
3f4210c13356 support for GHC: string literals;
wenzelm
parents:
diff changeset
    48
69211
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    49
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    50
(** antiquotations **)
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    51
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    52
(* theory data *)
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    53
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    54
structure Antiqs = Theory_Data
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    55
(
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    56
  type T = (Token.src -> Proof.context -> string) Name_Space.table;
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    57
  val empty : T = Name_Space.empty_table Markup.haskell_antiquotationN;
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    58
  val extend = I;
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    59
  fun merge tabs : T = Name_Space.merge_tables tabs;
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    60
);
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    61
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    62
val get_antiquotations = Antiqs.get o Proof_Context.theory_of;
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    63
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    64
fun antiquotation name scan body thy =
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    65
  let
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    66
    fun antiq src ctxt =
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    67
      let val (x, ctxt') = Token.syntax scan src ctxt
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    68
      in body {context = ctxt', source = src, argument = x} end;
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    69
  in thy |> Antiqs.map (Name_Space.define (Context.Theory thy) true (name, antiq) #> #2) end;
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    70
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    71
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    72
(* read source and expand antiquotations *)
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    73
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    74
val scan_antiq =
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    75
  Antiquote.scan_control >> Antiquote.Control ||
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    76
  Scan.one (Symbol.not_eof o Symbol_Pos.symbol) >> (Antiquote.Text o Symbol_Pos.symbol);
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    77
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    78
fun read_source ctxt source =
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    79
  let
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    80
    val _ =
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    81
      Context_Position.report ctxt (Input.pos_of source)
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    82
        (Markup.language_haskell (Input.is_delimited source));
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    83
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    84
    val (input, _) =
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    85
      Input.source_explode source
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    86
      |> Scan.error (Scan.finite Symbol_Pos.stopper (Scan.repeat scan_antiq));
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    87
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    88
    val _ = Context_Position.reports ctxt (Antiquote.antiq_reports input);
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    89
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    90
    fun expand antiq =
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    91
      (case antiq of
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    92
        Antiquote.Text s => s
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    93
      | Antiquote.Control {name, body, ...} =>
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    94
          let
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    95
            val src = Token.make_src name (if null body then [] else [Token.read_cartouche body]);
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    96
            val (src', f) = Token.check_src ctxt get_antiquotations src;
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    97
          in f src' ctxt end
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    98
      | Antiquote.Antiq antiq => error ("Bad antiquotation " ^ Position.here (#1 (#range antiq))));
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
    99
  in implode (map expand input) end;
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
   100
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
   101
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
   102
(* concrete antiquotation: ML expression as Haskell string literal *)
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
   103
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
   104
structure Local_Data = Proof_Data
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
   105
(
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
   106
  type T = string option;
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
   107
  fun init _ = NONE;
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
   108
);
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
   109
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
   110
val set_result = Local_Data.put o SOME;
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
   111
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
   112
fun the_result ctxt =
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
   113
  (case Local_Data.get ctxt of
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
   114
    SOME s => s
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
   115
  | NONE => raise Fail "No result");
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
   116
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
   117
val _ =
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
   118
  Theory.setup
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
   119
    (antiquotation \<^binding>\<open>cartouche\<close> (Scan.lift Args.cartouche_input)
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
   120
      (fn {context = ctxt, argument, ...} =>
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
   121
        ctxt |> Context.proof_map
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
   122
          (ML_Context.expression ("result", Position.thread_data ())
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
   123
            "string" "Context.map_proof (GHC.set_result result)"
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
   124
            (ML_Lex.read_source argument))
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
   125
        |> the_result |> print_string));
7062639cfdaa added GHC.read_source: read Haskell source text with antiquotations;
wenzelm
parents: 69209
diff changeset
   126
69209
3f4210c13356 support for GHC: string literals;
wenzelm
parents:
diff changeset
   127
end;