Library Apps.ExamplePwMgr

Require Import Ascii FunctionApp List Program.Basics StorageClient String System.
Require Import FunctionAppLemmas FunctionAppTactics.
Import ListNotations.
Open Scope string_scope.

Section ui.

  Inductive uiInput :=
  | uiConsoleIn : stringuiInput
  | uiDecrypted : stringuiInput.

  Context (world : Type).
  Context (consoleOut : stringaction world).
  Context (encrypt : stringaction world).

  Fixpoint split (sep : ascii) (s : string) : list string :=
    match s with
      | EmptyStringnil
      | String c s'
        if ascii_dec c sep then EmptyString :: split sep s'
        else match split sep s' with
               | nil[String c EmptyString]
               | w :: wsString c w :: ws
             end
    end.

  Definition newline := "010"%char.

  Definition dump (pws : list (string × string)) : string :=
    fold_right append ""%string
               (map (fun p ⇒ (fst p ++ " " ++ snd p ++ String newline "")%string) pws).

  Fixpoint load (s : string) : list (string × string) :=
    flat_map (fun lmatch split " " l with
                         | account :: password :: nil[(account, password)]
                         | _nil
                       end)
             (split newline s).

  Definition list_dec {A} (ls : list A) : { p | ls = fst p :: snd p } + { ls = nil }.
    destruct ls.
    - right; eauto.
    - left.
       (a, ls).
      eauto.
  Defined.

  Definition uiLoopBody (uiLoop : list (string × string) → process uiInput world) (pws : list (string × string))
  : uiInputaction world × process uiInput world
    := (fun i
          match i with
            | uiConsoleIn s
              match split " " s with
                | comm :: ls
                  match string_dec comm "get", ls with
                    | left _, account :: nil
                      match
                        find (fun pif string_dec account (fst p)
                                       then true else false) pws
                      with
                        | None
                          (consoleOut "account not found", uiLoop pws)
                        | Some (_, password)
                          (consoleOut password, uiLoop pws)
                      end
                    | _, _
                      match string_dec comm "set", ls with
                        | left _, account :: password :: nil
                          let pws' :=
                              (account, password)
                                :: filter (fun pif string_dec account (fst p)
                                                    then false else true) pws
                          in (encrypt (dump pws'), uiLoop pws')

                        | _, _
                          (consoleOut "unrecognized command", uiLoop pws)
                      end
                  end
                | _(consoleOut "unrecognized command", uiLoop pws)
              end
            | uiDecrypted s
              (id, uiLoop (load s))
          end).

  CoFixpoint uiLoop (pws : list (string × string)) :=
    Step (uiLoopBody uiLoop pws).

  Definition ui := uiLoop nil.

End ui.

Section net.

  Context {world0 : Type}.

  Inductive netInput :=
  | netReceived : abortable world0 (option string + httpResponse) → netInput
  | netEncrypted : stringnetInput.

  Context (world : Type).
  Context (httpPOST : stringlist (string × string) → (abortable world0 httpResponsenetInput) → action world).
  Context (decrypt : stringaction world).

  Definition storageId := "foo".

  CoFixpoint net :=
    Step (fun i
            match i with
              | netReceived (started abort) ⇒ (id, net)
              | netReceived (done (inl (Some s))) ⇒ (decrypt s, net)
              | netReceived (done (inl None)) ⇒ (id, net)
              | netReceived (done (inr _)) ⇒ (id, net)
              | netEncrypted s
                (storageSet httpPOST storageId "" s netReceived,
                 net)
            end).

End net.

Arguments netInput : clear implicits.
Arguments net : clear implicits.

Definition getStep {input output} (p : process input output) :=
  match p with
    | Step ff
  end.

Local Open Scope type_scope.

Section pwMgr.

  Context (world : Type).

  Inductive input :=
  | pwMgrConsoleIn : stringinput
  | pwMgrNetInput : netInput worldinput
  | pwMgrEncrypt : stringinput
  | pwMgrDecrypt : stringinput.

  Context (sys : systemActions input world).

  Definition pwMgrLoopBody pwMgrLoop ui net : inputaction (stackWorld input world) × stackProcess input world :=
    fun i
      match i with
        | pwMgrEncrypt s
          
          let (a, net') := getStep net (netEncrypted s) in (a, pwMgrLoop ui net')
        | pwMgrDecrypt s
          
          let (a, ui') := getStep ui (uiDecrypted s) in (a, pwMgrLoop ui' net)
        | pwMgrConsoleIn s
          let (a, ui') := getStep ui (uiConsoleIn s) in
          (fun wa (stackLift (consoleIn sys pwMgrConsoleIn) w), pwMgrLoop ui' net)
        | pwMgrNetInput i'
          let (a, net') := getStep net i' in (a, pwMgrLoop ui net')
      end.

  CoFixpoint pwMgrLoop ui net : stackProcess input world :=
    Step (pwMgrLoopBody pwMgrLoop ui net).

  Definition
    wrap_ui
    (ui :
        {world'},
         (stringaction world') →
         (stringaction world') →
         process uiInput world') :=
    ui
      (fun sstackLift (consoleOut sys s))
      (fun sstackPush (pwMgrEncrypt s)).

  Definition
    wrap_net
    (net :
        {world0 world'},
         (stringlist (string × string) → (abortable world0 httpResponsenetInput world0) → action world') →
         (stringaction world') →
         process (netInput world0) world') :=
    net
      (fun uri data cbstackLift (httpPOST sys uri data (fun rpwMgrNetInput (cb r))))
      (fun sstackPush (pwMgrDecrypt s)).

  Definition
    mkPwMgrStack ui net :
    stackProcess input world :=
    pwMgrLoop (wrap_ui ui) (wrap_net net).

  Definition pwMgrStack := mkPwMgrStack ui net.

  Lemma pwMgrLoop_eta ui net
  : pwMgrLoop ui net = Step (pwMgrLoopBody pwMgrLoop ui net).
  Proof.
    rewrite stackProcess_eta at 1; reflexivity.
  Qed.

  CoFixpoint pwMgrGood' :
     pws, emptiesStackForever
      (Step (pwMgrLoopBody pwMgrLoop (wrap_ui (fun world uiConsoleOut uiEncryptuiLoop world uiConsoleOut uiEncrypt pws)) (wrap_net net))).
  Proof.
    intro; constructor.
    let tac := (idtac;
                match goal with
                  | [ |- appcontext[match split ?a ?b with __ end] ] ⇒ destruct (split a b)
                  | [ |- appcontext[match string_dec ?s0 ?s1 with __ end] ] ⇒ destruct (string_dec s0 s1)
                  | [ |- appcontext[match ?l with nil_ | __ end] ] ⇒ destruct l
                  | [ |- appcontext[match find ?f ?ls with __ end] ] ⇒ destruct (find f ls)
                  | [ |- appcontext[match ?x with (_, _)_ end] ] ⇒ rewrite (@surjective_pairing _ _ x)
                  | [ |- appcontext[match ?a with (netReceived _) ⇒ _ | __ end] ] ⇒ destruct a
                  | [ |- appcontext[match ?a with (inl _) ⇒ _ | __ end] ] ⇒ destruct a
                  | [ |- appcontext[match ?a with None_ | __ end] ] ⇒ destruct a
                  | [ |- appcontext[match ?a with (done _) ⇒ _ | __ end] ] ⇒ destruct a
                  | _progress unfold storageSet in ×
                end) in
    emptiesStackForever_t pwMgrGood' input (@pwMgrLoop_eta) (@pwMgrLoop) tac.

    - subst; discriminate.
    - subst.
      constructor.
      constructor.
      Grab Existential Variables.
      eauto.
  Qed.

  Theorem pwMgrGood pws :
    emptiesStackForever
      (mkPwMgrStack
         (fun world uiConsoleOut uiEncryptuiLoop world uiConsoleOut uiEncrypt pws)
         net).
  Proof.
    unfold mkPwMgrStack.
    rewrite pwMgrLoop_eta.
    eapply pwMgrGood'.
  Qed.

  Definition proc := (consoleIn sys pwMgrConsoleIn, runStackProcess pwMgrStack (pwMgrGood nil)).

End pwMgr.

Require Import ExtrOcamlBasic ExtrOcamlString.
Extraction "ExamplePwMgr" proc.