LogrelCC.examples.refinement.server.server_nat

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

(*
let handler1 :  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 sum ->
*          writer (inr (sum + n)); writer (inl (associate tb (sum + n))); abort
*      end
*    | (None, n) ->
*      writer (inr n);
*      let cid = associate tb n
*      in writer (inl cid)
*)


(* server_nat_main below corresponds to the lines starred above. *)

(* Note the right to left nature of evaluation! *)

Definition server_nat_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 *)
          (LetIn
             (BinOp lang_base.Add (Var 0) (Snd (Var 2)))
             (Seq
                (App writer.[ren (+4)] (InjR (Var 0)))
                (Seq
                   (App writer.[ren (+4)] (InjL (App assoc.[ren (+4)] (Var 0))))
                   (Throw Unit (Cont []))
                )
             )
          )
       )
       (Seq
          (App writer.[ren (+2)] (InjR (Snd (Var 1))))
          (Seq
             (App writer.[ren (+2)] (InjL (App assoc.[ren (+2)] (Snd (Var 1)))))
             (Throw Unit (Cont []))
          )
       )
    ).

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


Lemma server_nat_main_eq reader writer get assoc :
 server_nat_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 *)
          (LetIn
             (BinOp lang_base.Add (Var 0) (Snd (Var 2)))
             (Seq
                (App writer.[ren (+4)] (InjR (Var 0)))
                (Seq
                   (App writer.[ren (+4)] (InjL (App assoc.[ren (+4)] (Var 0))))
                   (Throw Unit (Cont []))
                )
             )
          )
       )
       (Seq
          (App writer.[ren (+2)] (InjR (Snd (Var 1))))
          (Seq
             (App writer.[ren (+2)] (InjL (App assoc.[ren (+2)] (Snd (Var 1)))))
             (Throw Unit (Cont []))
          )
       )
    ).
Proof. trivial. Qed.

Typeclasses Opaque server_nat_main.
Global Opaque server_nat_main.

Lemma server_nat_main_subst f reader writer get assoc :
  (server_nat_main reader writer get assoc).[f] =
  server_nat_main reader.[f] writer.[f] get.[f] assoc.[f].
Proof. rewrite !server_nat_main_eq. by asimpl. Qed.

Hint Rewrite server_nat_main_subst : autosubst.

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


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

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