Library Io.System.Extraction
Extraction to OCaml. New primitives are defined in `extraction/utils.ml`.
Require Import Coq.Lists.List.
Require Import Coq.ZArith.ZArith.
Require Import ExtrOcamlBasic.
Require Import ExtrOcamlBigIntConv.
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.
Require Import Coq.ZArith.ZArith.
Require Import ExtrOcamlBasic.
Require Import ExtrOcamlBigIntConv.
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 with the OCaml strings.
The OCaml `string` type.
Export to an OCaml string.
Import an OCaml string.
Parameter to_lstring : t → LString.t.
Extract Constant to_lstring ⇒ "IoSystem.String.to_lstring".
End String.
Extract Constant to_lstring ⇒ "IoSystem.String.to_lstring".
End String.
Interface to the Big_int library.
The OCaml's `bigint` type.
Export to a `Z`.
Interface to the Sys library.
The command line arguments of the program.
Interface to the Lwt library.
The `Lwt.t` type.
Return.
Bind.
Run.
List the files of a directory.
Parameter list_files : String.t → t (option (list String.t)).
Extract Constant list_files ⇒ "IoSystem.list_files".
Extract Constant list_files ⇒ "IoSystem.list_files".
The the content of a file.
Parameter read_file : String.t → t (option String.t).
Extract Constant read_file ⇒ "IoSystem.read_file".
Extract Constant read_file ⇒ "IoSystem.read_file".
Update (or create) a file with some content.
Parameter write_file : String.t → String.t → t bool.
Extract Constant write_file ⇒ "IoSystem.write_file".
Extract Constant write_file ⇒ "IoSystem.write_file".
Delete a file.
Run a command.
Print a message on the standard output.
Read a line on the standard input.
Parameter read_line : unit → t (option String.t).
Extract Constant read_line ⇒ "IoSystem.read_line".
End Lwt.
Extract Constant read_line ⇒ "IoSystem.read_line".
End Lwt.
Evaluate a command using Lwt.
Definition eval_command (c : Effects.command System.effects)
: Lwt.t (Effects.answer System.effects c) :=
match c return (Lwt.t (Effects.answer System.effects 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 command ⇒ Lwt.system (String.of_lstring command)
| 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.
: Lwt.t (Effects.answer System.effects c) :=
match c return (Lwt.t (Effects.answer System.effects 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 command ⇒ Lwt.system (String.of_lstring command)
| 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.effects A) : Lwt.t A :=
match x with
| C.Ret _ x ⇒ Lwt.ret x
| C.Call command ⇒ eval_command command
| C.Let _ _ x f ⇒ Lwt.bind (eval x) (fun x ⇒ eval (f x))
end.
match x with
| C.Ret _ x ⇒ Lwt.ret x
| C.Call command ⇒ eval_command command
| C.Let _ _ x f ⇒ Lwt.bind (eval x) (fun x ⇒ eval (f x))
end.
Run the main function.
Definition run (main : list LString.t → C.t System.effects unit) : unit :=
let argv := List.map String.to_lstring Sys.argv in
Extraction.Lwt.run (Extraction.eval (main argv)).
let argv := List.map String.to_lstring Sys.argv in
Extraction.Lwt.run (Extraction.eval (main argv)).