Library Apps.TrustedServerSyncBox

A box to mirror a server over untrusted channels using encryption while preventing timing side channels (part of TCB)

Require Import Coq.Program.Basics Coq.NArith.NArith Coq.Lists.List.
Require Import FunctionApp EncryptionInterface TrustedTickBox TrustedEncryptBox TrustedDecryptBox.
Require Import Common.

Local Open Scope list_scope.
Local Open Scope program_scope.

Set Implicit Arguments.

Summary

We implement a box that presents the interface of a server storing unencrypted data. It handles "set" and "get eventually" requests, and raises "value changed" and "remote corrupted".
We assume the remote server has a "compare and set" event; we use this to be robust to update messages lost over the network. It is the client's responsibility to set up the tick-box intervals so that we get an updated server value sufficiently close to our sending that CAS succeeds.
We maintain the following internal state:
  • remoteStateE - state we think the remote server has right now (encrypted)
  • localStateD - state we want the remote server to have (decrypted)
We implement the following events:
Client Events:
  • ssbClientGetUpdate - When the client requests an update from the server, we request an update from our remote server, hidden behind a tick box.
  • ssbClientSet (newD : rawDataT) - When the client requests a SET, we set localStateD to newD, and schedule a remote update.
Server Events:
  • ssbServerGotUpdate (newE : encryptedDataT) - When the server provides us with an update, we decrypt it to newD. If we fail, we inform the client that the remote server got messed up. If we succeed, and it is different from localStateD, we set localStateD to newD, and push an update to the client. In all cases, we set remoteStateE to newE.
Timed events:
  • When it's time to do a server CAS, we compute an encryption localStateE of localStateD. We then tell the server that if it's state is remoteStateE, it should replace it with localStateE. The server is expected to respond with ssbServerGotUpdate with it's updated state on both failure and success.

Module TrustedServerSyncBox (DataTypes : EncryptionDataTypes) (Algorithm : EncryptionAlgorithm DataTypes).
  Import DataTypes.

  Module Import TEB := TrustedEncryptBox DataTypes Algorithm.
  Module Import TDB := TrustedDecryptBox DataTypes Algorithm.

  Section trustedServerSyncBox.
    Variable d_eqb : rawDataTrawDataTbool.

    Record ServerSyncBoxPreState
      := { localStateD : rawDataT;
           remoteStateE : option encryptedDataT }.

    Record ServerSyncBoxState :=
      {
        ssbState :> ServerSyncBoxPreState;
        ssbGetUpdateState : TickBoxState unit;
        ssbCASState : TickBoxState encryptedDataT;
        ssbEncryptState : EncryptBoxState;
        ssbDecryptState : DecryptBoxState
      }.


    Definition set_ssbState (st : ServerSyncBoxState) (v : ServerSyncBoxPreState)
      := {| ssbState := v;
            ssbGetUpdateState := st.(ssbGetUpdateState);
            ssbCASState := st.(ssbCASState);
            ssbEncryptState := st.(ssbEncryptState);
            ssbDecryptState := st.(ssbDecryptState) |}.

    Definition set_ssbGetUpdateState (st : ServerSyncBoxState) (v : TickBoxState unit)
      := {| ssbState := st.(ssbState);
            ssbGetUpdateState := v;
            ssbCASState := st.(ssbCASState);
            ssbEncryptState := st.(ssbEncryptState);
            ssbDecryptState := st.(ssbDecryptState) |}.

    Definition set_ssbCASState (st : ServerSyncBoxState) (v : TickBoxState encryptedDataT)
      := {| ssbState := st.(ssbState);
            ssbGetUpdateState := st.(ssbGetUpdateState);
            ssbCASState := v;
            ssbEncryptState := st.(ssbEncryptState);
            ssbDecryptState := st.(ssbDecryptState) |}.

    Definition set_ssbEncryptState (st : ServerSyncBoxState) (v : EncryptBoxState)
      := {| ssbState := st.(ssbState);
            ssbGetUpdateState := st.(ssbGetUpdateState);
            ssbCASState := st.(ssbCASState);
            ssbEncryptState := v;
            ssbDecryptState := st.(ssbDecryptState) |}.

    Definition set_ssbDecryptState (st : ServerSyncBoxState) (v : DecryptBoxState)
      := {| ssbState := st.(ssbState);
            ssbGetUpdateState := st.(ssbGetUpdateState);
            ssbCASState := st.(ssbCASState);
            ssbEncryptState := st.(ssbEncryptState);
            ssbDecryptState := v |}.

    Inductive ssbConfigInput :=
    | ssbGetUpdateConfig (_ : tbConfigInput)
    | ssbCASConfig (_ : tbConfigInput)
    | ssbSetMasterKey (key : masterKeyT).

    Inductive ssbEventInput :=
    | ssbClientGetUpdate
    | ssbClientSet (newD : rawDataT)
    | ssbServerGotUpdate (newE : encryptedDataT)
    | ssbSystemRandomness (randomness : systemRandomnessT) (tag : rawDataT)
    | ssbTick (_ : N).

    Definition ssbInput := (ssbConfigInput + ssbEventInput)%type.

    Inductive ssbWarningOutput :=
    | ssbGetUpdateWarning (_ : tbWarningOutput unit)
    | ssbCASWarning (_ : tbWarningOutput encryptedDataT)
    | ssbEncryptError (_ : ebErrorOutput)
    | ssbDecryptError (_ : dbErrorOutput unit)
    | ssbWarningInvalidTransition (ev : ssbInput) (st : ServerSyncBoxState)
    | ssbWarningPushBeforePull.

    Inductive ssbEventOutput :=
    | ssbClientGotUpdate (data : rawDataT)
    | ssbServerGetUpdate
    | ssbServerCAS (curE newE : encryptedDataT)
    | ssbRequestSystemRandomness (howMuch : systemRandomnessHintT) (tag : rawDataT)
    | ssbSleepFor (ticks : N).

    Definition ssbOutput := (ssbWarningOutput + ssbEventOutput)%type.

    Variable initRawData : rawDataT.
    Context (world : Type)
            (handle : ssbOutputaction world).

    Definition initState : ServerSyncBoxState :=
      {|
        ssbState := {| localStateD := initRawData;
                       remoteStateE := None |};
        ssbGetUpdateState := TrustedTickBox.initState _;
        ssbCASState := TrustedTickBox.initState _;
        ssbEncryptState := TEB.initState;
        ssbDecryptState := TDB.initState
      |}.

    Fixpoint isTickBoxSimpleResult {T} (res : list (tbOutput T)) : bool
      := match res with
           | niltrue
           | inl _::res'isTickBoxSimpleResult res'
           | inr (tbSleepFor _)::res'isTickBoxSimpleResult res'
           | _false
         end.

    Fixpoint handle_simpleTickBox {T} (warn : tbWarningOutput TssbWarningOutput) (res : list (tbOutput T)) {struct res}
    : isTickBoxSimpleResult res = truelist ssbOutput
      := match res as res return isTickBoxSimpleResult res = true_ with
           | nilfun _nil
           | inl warning::res'fun Hinl (warn warning)::(@handle_simpleTickBox _ warn res' H)
           | inr (tbSleepFor ticks)::res'fun Hinr (ssbSleepFor ticks)::(@handle_simpleTickBox _ warn res' H)
           | _fun H : false = truematch Bool.diff_false_true H with end
         end.

    Local Ltac handle_isSimple :=
      subst_body; clear;
      abstract (
          repeat match goal with
                   | _reflexivity
                   | [ |- appcontext[match ?x with __ end] ] ⇒ (atomic x; destruct x)
                   | [ |- isTickBoxSimpleResult ?x = _ ] ⇒ let x' := (eval hnf in x) in progress change x with x'
                   | [ |- appcontext[match ?x with __ end] ] ⇒ let x' := (eval hnf in x) in progress change x with x'
                 end
        ).

    Definition handle_ssbGetUpdate'
               (st : ServerSyncBoxState)
    : tbOutput unit × TickBoxState unitlist ssbOutput × ServerSyncBoxState.
    Proof.
      refine (fun i
                (match fst i with
                   | inl warning
                     ⇒ inl (ssbGetUpdateWarning warning)::nil
                   | inr tbRequestDataUpdate
                     ⇒ let upd := tickBoxLoopPreBody
                                     (st.(ssbGetUpdateState))
                                     (inr (tbValueReady tt)) in
                        handle_simpleTickBox ssbGetUpdateWarning (fst upd) _
                   | inr (tbPublishUpdate val)
                     ⇒ inr ssbServerGetUpdate::nil
                   | inr (tbSleepFor ticks)
                     ⇒ inr (ssbSleepFor ticks)::nil
                 end,
                 set_ssbGetUpdateState st (snd i)));
      handle_isSimple.
    Defined.

    Definition fold_handler {outT stT}
               (get : ServerSyncBoxStatestT)
               (set : ServerSyncBoxStatestTServerSyncBoxState)
               (handle : ServerSyncBoxStateoutT × stTlist ssbOutput × ServerSyncBoxState)
               (st : ServerSyncBoxState)
    : list outT × stTlist ssbOutput × ServerSyncBoxState
      := fun outs_st
           fold_left (fun ls_st out
                        let ls'_st' := handle (snd ls_st) (out, get (snd ls_st)) in
                        (fst ls_st ++ fst ls'_st', snd ls'_st'))
                     (fst outs_st)
                     (nil, set st (snd outs_st)).

    Definition handle_ssbGetUpdate
      := fold_handler ssbGetUpdateState set_ssbGetUpdateState handle_ssbGetUpdate'.

    Definition handle_ssbEncrypt
               (st : ServerSyncBoxState)
    : option (ebOutput unit) × EncryptBoxStatelist ssbOutput × ServerSyncBoxState.
    Proof.
      refine (fun i
                let st' := set_ssbEncryptState st (snd i) in
                match fst i with

                   | Some (inl warning)
                     ⇒ ((inl (ssbEncryptError warning))::nil, st')

                   | Some (inr (ebRequestSystemRandomness howMuch tag))
                     ⇒ ((inr (ssbRequestSystemRandomness howMuch (fst tag)))::nil, st')

                   | Some (inr (ebEncrypted newE _))
                     ⇒ let upd := tickBoxLoopPreBody
                                     (st'.(ssbCASState))
                                     (inr (tbValueReady newE)) in
                        (handle_simpleTickBox ssbCASWarning (fst upd) _,
                         set_ssbCASState st' (snd upd))

                   | None(nil, st')
                 end);
      handle_isSimple.
    Defined.

    Definition handle_ssbCAS'
               (st : ServerSyncBoxState)
    : tbOutput encryptedDataT × TickBoxState encryptedDataTlist ssbOutput × ServerSyncBoxState.
    Proof.
      refine (fun i
                let st' := set_ssbCASState st (snd i) in
                match fst i with
                  | inl warning
                    ⇒ (inl (ssbCASWarning warning)::nil, st')

                  | inr (tbPublishUpdate val)
                    ⇒ (match st'.(remoteStateE) with
                          | Some curE
                            ⇒ inr (ssbServerCAS curE val)
                          | Noneinl ssbWarningPushBeforePull
                        end::nil,
                        st')

                  | inr tbRequestDataUpdate
                    ⇒ let encResult := encryptBoxLoopPreBody
                                          (st'.(ssbEncryptState))
                                          (ebEncrypt (st'.(localStateD)) tt) in
                       handle_ssbEncrypt st' encResult

                  | inr (tbSleepFor ticks)
                    ⇒ (inr (ssbSleepFor ticks)::nil, st')
                end).
    Defined.

    Definition handle_ssbCAS
      := fold_handler ssbCASState set_ssbCASState handle_ssbCAS'.

    Definition handle_ssbDecrypt
               (st : ServerSyncBoxState)
    : option (dbOutput unit) × DecryptBoxStatelist ssbOutput × ServerSyncBoxState.
    Proof.
      refine (fun i
                let st' := set_ssbDecryptState st (snd i) in
                match fst i with

                  | Some (inl warning)
                    ⇒ (inl (ssbDecryptError warning)::nil, st')

                  | Some (inr (dbDecrypted dataD _))
                    ⇒ ((if d_eqb dataD st'.(localStateD)
                         then nil
                         else (inr (ssbClientGotUpdate dataD))::nil),
                        set_ssbState st' {| localStateD := dataD;
                                            remoteStateE := st'.(remoteStateE) |})

                  | None(nil, st')

                end).
    Defined.

    Definition serverSyncBoxLoopPreBody
               (st : ServerSyncBoxState)
    : ssbInputlist ssbOutput × ServerSyncBoxState.
    Proof.
      refine (fun i
                match i with

                  | inl (ssbGetUpdateConfig i')
                    ⇒ handle_ssbGetUpdate
                         st
                         (tickBoxLoopPreBody
                            (st.(ssbGetUpdateState))
                            (inl i'))

                  | inl (ssbCASConfig i')
                    ⇒ handle_ssbCAS
                         st
                         (tickBoxLoopPreBody
                            (st.(ssbCASState))
                            (inl i'))

                  | inl (ssbSetMasterKey newKey)
                    ⇒ let (ls0, st0) := handle_ssbEncrypt
                                           st
                                           (encryptBoxLoopPreBody
                                              (st.(ssbEncryptState))
                                              (ebSetMasterKey _ newKey)) in
                       let (ls1, st1) := handle_ssbDecrypt
                                           st0
                                           (decryptBoxLoopPreBody
                                              (st.(ssbDecryptState))
                                              (dbSetMasterKey _ newKey)) in
                       (ls0 ++ ls1, st1)

                  | inr (ssbTick n)
                    ⇒ let (ls0, st0) := handle_ssbGetUpdate
                                           st
                                           (tickBoxLoopPreBody
                                              (st.(ssbGetUpdateState))
                                              (inr (tbTick _ n))) in
                       let (ls1, st1) := handle_ssbCAS
                                           st0
                                           (tickBoxLoopPreBody
                                              (st0.(ssbCASState))
                                              (inr (tbTick _ n))) in
                       (ls0 ++ ls1, st1)

                  | inr ssbClientGetUpdate
                    ⇒ handle_ssbGetUpdate
                         st
                         (tickBoxLoopPreBody
                            (st.(ssbGetUpdateState))
                            (inr (tbNotifyChange _)))

                  | inr (ssbClientSet newD)
                    ⇒ let st' := set_ssbState
                                    st
                                    {| localStateD := newD;
                                       remoteStateE := st.(remoteStateE) |} in
                       handle_ssbCAS
                         st'
                         (tickBoxLoopPreBody
                            (st'.(ssbCASState))
                            (inr (tbNotifyChange _)))

                  | inr (ssbServerGotUpdate dataE)
                    ⇒ let st' := set_ssbState
                                    st
                                    {| localStateD := st.(localStateD);
                                       remoteStateE := Some dataE |} in
                       handle_ssbDecrypt
                         st'
                         (decryptBoxLoopPreBody
                            (st'.(ssbDecryptState))
                            (dbDecrypt dataE tt))

                  | inr (ssbSystemRandomness randomness tag)
                    ⇒ handle_ssbEncrypt
                         st
                         (encryptBoxLoopPreBody
                            (st.(ssbEncryptState))
                            (ebSystemRandomness randomness (tag, tt)))

                end).
    Defined.

    Definition serverSyncBoxLoopBody {T}
               (serverSyncBoxLoop : ServerSyncBoxStateT)
               (st : ServerSyncBoxState)
    : ssbInputaction world × T
      := fun ilet outs := fst (serverSyncBoxLoopPreBody st i) in
                  (fold_left compose (map handle outs) id,
                   serverSyncBoxLoop (snd (serverSyncBoxLoopPreBody st i))).

    CoFixpoint serverSyncBoxLoop (st : ServerSyncBoxState) :=
      Step (serverSyncBoxLoopBody serverSyncBoxLoop st).

    Definition serverSyncBox : process _ _ := serverSyncBoxLoop initState.
  End trustedServerSyncBox.
End TrustedServerSyncBox.