LogrelCC.examples.refinement.server.server_cont

From LogrelCC Require Export lang.
From LogrelCC.examples Require Import list_assoc.

(*
* let read_client tb writer = callcc (k, writer (inl (associate tb k));  abort)

let handler2 : ServerConnT ->  =
  let tb = newTable () in
  fun (cn : ServerConnT) ->
*    let (reader, writer) = cn in
*    match reader () with
*      (Some cid, n) ->
*      begin
*        match get tb cid with
*        | None -> () (* unknown resumption id! *)
*        | Some k -> throw (n, reader, writer) to k
*      end
*    | inr (None, n) ->
*      let rec loop (m, writer) =
*        writer (inr m);
*        let (v, writer) = read_client () in loop (m + v, writer)
*      in loop (n, writer)
*)


(* server_cont_main below corresponds to the lines starred above. We
have inlined read_client above as it appears only once in this code *)


Definition read_client assoc writer :=
  Callcc
    (Seq
       (LetIn
          (InjL (App assoc.[ren (+1)] (Var 0)))
          (App writer.[ren (+2)] (Var 0))
       )
       (Throw Unit (Cont []))
    ).

Definition server_cont_main reader writer get assoc :=
  LetIn
    (App reader Unit)
    (Case
       (Fst (Var 0))
       (*client passed in a resumption id *)
       (Case (App get.[ren (+2)] (Var 0))
             (* invalid resumption id *)
             Unit
             (* valid resumption id *)
             (Throw (Pair (Snd (Var 2)) writer.[ren (+3)]) (Var 0)))
       (*no session id passed by client*)
       (App
          (Rec
             (LetIn (* given value *)
                (Fst (Var 1))
                (LetIn (* writer *)
                   (Snd (Var 2))
                   (Seq
                      (App (Var 0) (InjR (Var 1)))
                      (LetIn
                         (read_client assoc.[ren (+6)] (Var 0))
                         (App (Var 3)
                              (Pair
                                 (BinOp lang_base.Add (Var 2) (Fst (Var 0)))
                                 (Snd (Var 0)))))
                   )
                )
             )
          )
          (Pair (Snd (Var 1)) writer.[ren (+2)])
       )
    ).

(*   App (Rec *)
(*          (Case *)
(*             (Fst (Var 1)) *)
(*             (*client passed in a resumption id *) *)
(*             (Case (App get.ren (+3) (Var 0)) *)
(*                   (* invalid resumption id *) *)
(*                   Unit *)
(*                   (* valid resumption id *) *)
(*                   (Throw (Pair (Snd (Var 3)) writer.ren (+4)) (Var 0))) *)
(*             (*no session id passed by client*) *)
(*             (App *)
(*                (Rec *)
(*                   (App *)
(*                      (Rec *)
(*                         (App *)
(*                            (Rec *)
(*                               (App *)
(*                                  (Rec *)
(*                                     (App *)
(*                                        (Var 6) *)
(*                                        (Pair *)
(*                                           (BinOp Add (Var 5) (Fst (Var 1))) *)
(*                                           (Snd (Var 1))))) *)
(*                                  (App *)
(*                                     (Rec (Callcc *)
(*                                             (App *)
(*                                                (Rec (Throw Unit (Cont ))) *)
(*                                                (App *)
(*                                                   (Rec *)
(*                                                      (App (Var 6) *)
(*                                                           (InjL (Var 1)))) *)
(*                                                   (* associate the current continuation *) *)
(*                                                   (* to a resumption id *) *)
(*                                                   (App assoc.ren (+12) (Var 0)))))) *)
(*                                     (* write back the current sum. *) *)
(*                                     (App (Var 1) (InjR (Var 3)))))) *)
(*                            (Snd (Var 3)))) *)
(*                      (Fst (Var 1)) *)
(*                )) *)
(*                (Pair (Snd (Var 2)) writer.ren (+3))) *)

(* )) *)
(*       (App reader Unit). *)

(* The following definitions and lemmas are mainly to battle
Autosubst's complexity by defining simplification lemmas by hand and
applying them. *)


Lemma server_cont_main_eq reader writer get assoc :
  server_cont_main reader writer get assoc =
  LetIn
    (App reader Unit)
    (Case
       (Fst (Var 0))
       (*client passed in a resumption id *)
       (Case (App get.[ren (+2)] (Var 0))
             (* invalid resumption id *)
             Unit
             (* valid resumption id *)
             (Throw (Pair (Snd (Var 2)) writer.[ren (+3)]) (Var 0)))
       (*no session id passed by client*)
       (App
          (Rec
             (LetIn (* given value *)
                (Fst (Var 1))
                (LetIn (* writer *)
                   (Snd (Var 2))
                   (Seq
                      (App (Var 0) (InjR (Var 1)))
                      (LetIn
                         (read_client assoc.[ren (+6)] (Var 0))
                         (App (Var 3)
                              (Pair
                                 (BinOp lang_base.Add (Var 2) (Fst (Var 0)))
                                 (Snd (Var 0)))))
                   )
                )
             )
          )
          (Pair (Snd (Var 1)) writer.[ren (+2)])
       )
    ).
Proof. trivial. Qed.

Typeclasses Opaque server_cont_main.
Global Opaque server_cont_main.

Lemma server_cont_main_subst f reader writer get assoc :
  (server_cont_main reader writer get assoc).[f] =
  server_cont_main reader.[f] writer.[f] get.[f] assoc.[f].
Proof. rewrite !server_cont_main_eq !/read_client. by asimpl. Qed.

Hint Rewrite server_cont_main_subst : autosubst.

(* This wraps server_cont_main above to complete the code (add the
parts that are not starred). *)


Definition server_cont :=
    LetIn
    (TApp make_association)
    (LetIn
       (Fst (Var 0))
       (LetIn
          (Snd (Var 1))
          (Lam
             (LetIn
                (Fst (Var 0))
                (LetIn
                   (Snd (Var 1))
                   (server_cont_main (Var 1) (Var 0) (Var 3) (Var 4))))
          )
       )
    ).

  (* App (Rec *)
  (*        (App *)
  (*           (Rec *)
  (*              (App *)
  (*                 (Rec *)
  (*                    (Rec *)
  (*                       (App *)
  (*                          (Rec *)
  (*                             (App *)
  (*                                (Rec (server_cont_main *)
  (*                                        (Var 3) (Var 1) (Var 7) (Var 9))) *)
  (*                                (Snd (Var 3)))) *)
  (*                          (Fst (Var 1)))) *)
  (*                 ) *)
  (*                 (Snd (Var 3)))) *)
  (*           (Fst (Var 1)))) *)
  (*     (TApp make_association). *)