Library Io.System.Extraction

Extraction to OCaml. New primitives are defined in `extraction/utils.ml`.
Require Import Coq.Lists.List.
Require Import Coq.PArith.PArith.
Require Import Coq.ZArith.ZArith.
Require Import ExtrOcamlBasic.
Require Import ExtrOcamlString.
Require Import ErrorHandlers.All.
Require Import FunctionNinjas.All.
Require Import ListString.All.
Require Import Io.All.
Require System.

Import ListNotations.
Local Open Scope type.

Interface to the Big_int library.
Module BigInt.
The OCaml's `big_int` type.
  Parameter t : Type.
  Extract Constant t ⇒ "Big_int.big_int".

  Parameter to_Z_aux : {Z positive : Type},
    t
    Z → (positiveZ) → (positiveZ) →
    positive → (positivepositive) → (positivepositive) →
    Z.
  Extract Constant to_Z_aux ⇒ "IoSystem.Big.to_Z_aux".

Export to a `Z`.
  Definition to_Z (big : t) : Z :=
    to_Z_aux big Z0 Zpos Zneg xH xO xI.
End BigInt.

Interface with the OCaml strings.
Module String.
The OCaml `string` type.
  Parameter t : Type.
  Extract Constant t ⇒ "string".

Export to an OCaml string.
  Parameter of_lstring : LString.tt.
  Extract Constant of_lstring ⇒ "IoSystem.String.of_lstring".

Import an OCaml string.
  Parameter to_lstring : tLString.t.
  Extract Constant to_lstring ⇒ "IoSystem.String.to_lstring".
End String.

Interface to the sum type.
Module Sum.
  Parameter t : TypeTypeType.
  Extract Constant t "'a" "'b" ⇒ "('a, 'b) IoSystem.Sum.t".

  Parameter destruct : {A B C : Type}, t A B → (AC) → (BC) → C.
  Extract Constant destruct ⇒ "IoSystem.Sum.destruct".

  Definition to_coq {A B : Type} (s : t A B) : A + B :=
    destruct s inl inr.
End Sum.

Interface to the Sys library.
Module Sys.
The command line arguments of the program.
  Parameter argv : list String.t.
  Extract Constant argv ⇒ "IoSystem.argv".
End Sys.

Interface to the Lwt library.
Module Lwt.
The `Lwt.t` type.
  Parameter t : TypeType.
  Extract Constant t "'a" ⇒ "'a Lwt.t".

Return.
  Parameter ret : {A : Type}, At A.
  Extract Constant ret ⇒ "Lwt.return".

Bind.
  Parameter bind : {A B : Type}, t A → (At B) → t B.
  Extract Constant bind ⇒ "Lwt.bind".

Join.
  Parameter join : {A B : Type}, t At Bt (A × B).
  Extract Constant join ⇒ "IoSystem.join".

First.
  Parameter first : {A B : Type}, t At Bt (Sum.t A B).
  Extract Constant first ⇒ "IoSystem.first".

Run.
  Parameter run : {A : Type}, t AA.
  Extract Constant run ⇒ "Lwt_main.run".

List the files of a directory.
  Parameter list_files : String.tt (option (list String.t)).
  Extract Constant list_files ⇒ "IoSystem.list_files".

The the content of a file.
  Parameter read_file : String.tt (option String.t).
  Extract Constant read_file ⇒ "IoSystem.read_file".

Update (or create) a file with some content.
  Parameter write_file : String.tString.tt bool.
  Extract Constant write_file ⇒ "IoSystem.write_file".

Delete a file.
  Parameter delete_file : String.tt bool.
  Extract Constant delete_file ⇒ "IoSystem.delete_file".

Run a command.
  Parameter system : String.tt (option bool).
  Extract Constant system ⇒ "IoSystem.system".

Run a command controlling the outputs.
  Parameter eval : list String.tt (option (BigInt.t × String.t × String.t)).
  Extract Constant eval ⇒ "IoSystem.eval".

Print a message on the standard output.
  Parameter print : String.tt bool.
  Extract Constant print ⇒ "IoSystem.print".

Read a line on the standard input.
  Parameter read_line : unitt (option String.t).
  Extract Constant read_line ⇒ "IoSystem.read_line".
End Lwt.

Evaluate a command using Lwt.
Definition eval_command (c : Effect.command System.effect)
  : Lwt.t (Effect.answer System.effect c) :=
  match c return (Lwt.t (Effect.answer System.effect c)) with
  | System.ListFiles folder
    Lwt.bind (Lwt.list_files @@ String.of_lstring folder) (fun files
    Lwt.ret @@ Option.bind files (fun files
    Some (List.map String.to_lstring files)))
  | System.ReadFile file_name
    Lwt.bind (Lwt.read_file @@ String.of_lstring file_name) (fun content
    Lwt.ret @@ option_map String.to_lstring content)
  | System.WriteFile file_name content
    let file_name := String.of_lstring file_name in
    let content := String.of_lstring content in
    Lwt.write_file file_name content
  | System.DeleteFile file_name
    Lwt.delete_file @@ String.of_lstring file_name
  | System.System commandLwt.system (String.of_lstring command)
  | System.Eval command
    let command := List.map String.of_lstring command in
    Lwt.bind (Lwt.eval command) (fun result
    Lwt.ret (result |> option_map (fun result
    match result with
    | (status, output, err)
      (BigInt.to_Z status, String.to_lstring output, String.to_lstring err)
    end)))
  | System.Print message
    let message := String.of_lstring message in
    Lwt.print message
  | System.ReadLine
    Lwt.bind (Lwt.read_line tt) (fun line
    Lwt.ret @@ option_map String.to_lstring line)
  end.

Evaluate an expression using Lwt.
Fixpoint eval {A : Type} (x : C.t System.effect A) : Lwt.t A.
  destruct x as [A x | command | A B x f | A B x y | A B x y].
  - exact (Lwt.ret x).
  - exact (eval_command command).
  - exact (Lwt.bind (eval _ x) (fun xeval _ (f x))).
  - exact (Lwt.join (eval _ x) (eval _ y)).
  - exact (
      Lwt.bind (Lwt.first (eval _ x) (eval _ y)) (fun s
      Lwt.ret @@ Sum.to_coq s)).
Defined.

Run the main function.
Definition run (main : list LString.tC.t System.effect unit) : unit :=
  let argv := List.map String.to_lstring Sys.argv in
  Lwt.run (Extraction.eval (main argv)).