libisabelle-protocol/libisabelle/protocol/Protocol.thy
author Walther Neuper <wneuper@ist.tugraz.at>
Tue, 04 Sep 2018 14:50:30 +0200
changeset 59469 5c56f14bea53
parent 59340 097347b8910e
permissions -rw-r--r--
Isabelle2017->18: add libisabelle, PROBLEM with session management:

/usr/local/isabisac/src/Tools/jEdit/dist/jars/Isabelle-jEdit.jar:
Cannot start:
*** Duplicate theory "Protocol.Protocol" vs. "/usr/local/isabisac/libisabelle-protocol/libisabelle/protocol/Protocol.thy"
wneuper@59340
     1
theory Protocol
wneuper@59340
     2
imports Codec_Class
wneuper@59340
     3
keywords "operation_setup" :: thy_decl % "ML"
wneuper@59340
     4
begin
wneuper@59340
     5
wneuper@59340
     6
ML\<open>
wneuper@59340
     7
val _ =
wneuper@59340
     8
  let
wneuper@59340
     9
    open Libisabelle_Protocol
wneuper@59340
    10
    fun operation_setup_cmd name source (flags as {auto, ...}) ctxt =
wneuper@59340
    11
      let
wneuper@59340
    12
        fun eval enclose =
wneuper@59340
    13
          ML_Context.eval_in (SOME ctxt) ML_Compiler.flags (Input.pos_of source)
wneuper@59340
    14
            (ML_Lex.read ("Libisabelle_Protocol.add_operation " ^ ML_Syntax.print_string name ^ "(") @
wneuper@59340
    15
              enclose (ML_Lex.read_source false source) @
wneuper@59340
    16
              ML_Lex.read ")" @
wneuper@59340
    17
              ML_Lex.read (print_flags flags))
wneuper@59340
    18
      in
wneuper@59340
    19
        if auto then
wneuper@59340
    20
          let
wneuper@59340
    21
            (* FIXME breaks antiquotations *)
wneuper@59340
    22
            val ML_Types.Fun (arg, res) = ML_Types.ml_type_of ctxt (Input.source_content source)
wneuper@59340
    23
            val arg_codec = Classy.resolve @{ML.class codec} arg (Context.Proof ctxt)
wneuper@59340
    24
            val res_codec = Classy.resolve @{ML.class codec} res (Context.Proof ctxt)
wneuper@59340
    25
            fun enclose toks =
wneuper@59340
    26
              ML_Lex.read "{from_lib=" @ ML_Lex.read arg_codec @ ML_Lex.read "," @
wneuper@59340
    27
              ML_Lex.read "to_lib=" @ ML_Lex.read res_codec @ ML_Lex.read "," @
wneuper@59340
    28
              ML_Lex.read "action=" @ toks @ ML_Lex.read "}"
wneuper@59340
    29
          in
wneuper@59340
    30
            eval enclose
wneuper@59340
    31
          end
wneuper@59340
    32
        else
wneuper@59340
    33
          eval I
wneuper@59340
    34
      end
wneuper@59340
    35
    val parse_flag =
wneuper@59340
    36
      (Parse.reserved "sequential" || Parse.reserved "bracket" || Parse.reserved "auto") >>
wneuper@59340
    37
        (fn flag => join_flags
wneuper@59340
    38
           {sequential = flag = "sequential",
wneuper@59340
    39
            bracket = flag = "bracket",
wneuper@59340
    40
            auto = flag = "auto"})
wneuper@59340
    41
    val parse_flags =
wneuper@59340
    42
      Parse.list parse_flag >> (fn fs => fold (curry op o) fs I)
wneuper@59340
    43
    val parse_cmd =
wneuper@59340
    44
      Scan.optional (Args.parens parse_flags) I --
wneuper@59340
    45
      Parse.name --
wneuper@59340
    46
      Parse.!!! (@{keyword "="} |-- Parse.ML_source)
wneuper@59340
    47
  in
wneuper@59340
    48
    Outer_Syntax.command @{command_keyword "operation_setup"} "define protocol operation in ML"
wneuper@59340
    49
      (parse_cmd >> (fn ((flags, name), txt) =>
wneuper@59340
    50
        Toplevel.keep (Toplevel.context_of #> operation_setup_cmd name txt (flags default_flags))))
wneuper@59340
    51
  end
wneuper@59340
    52
\<close>
wneuper@59340
    53
wneuper@59340
    54
end