Check nat. (* STLC *) Inductive type : Type := | Bool: type | Arrow: type -> type -> type. Section stlc. (*Variable V : type -> Type. *) Context {V : type -> Type}. Inductive term : type -> Type := | Var : forall t:type, V t -> term t | Tru : term Bool | Fals : term Bool | App : forall (t1 t2 : type), term (Arrow t1 t2) -> term t1 -> term t2 | Abs : forall (t1 t2 : type), (V t1 -> term t2) -> term (Arrow t1 t2). End stlc. Fixpoint typeDenote (t : type) : Type := match t with | Bool => bool | Arrow t1 t2 => (typeDenote t1) -> (typeDenote t2) end. Fixpoint termDenote {t : type} (e : term t) : typeDenote t := (*match e in (term _ t) return (typeDenote t) with*) match e with (* _ is for the t *) | Var _ v => v | Tru => true | Fals => false | App _ _ e1 e2 => (termDenote e1) (termDenote e2) | Abs _ _ f => (* XXX: why was the fun abstr needed here?? *) fun x => termDenote (f x) end. (* CPS *) Inductive ptype : Type := | PBool : ptype | PCont : ptype -> ptype (* Continuation type *) | PUnit : ptype (* Useful for PCont?? *) | PProd : ptype -> ptype -> ptype. Fixpoint ptypeDenote (t : ptype) : Type := match t with | PBool => bool | PCont t' => ptypeDenote t' -> bool (* τ → 0 *) | PUnit => unit | PProd t1 t2 => (ptypeDenote t1) * (ptypeDenote t2) end. Section cpsterm. (*Variable V : ptype -> Type. Variable res : ptype.*) Context {V : ptype -> Type} {res : ptype}. Inductive pterm : Type := | PHalt : V res -> pterm | PApp : forall (t:ptype), V (PCont t) -> V t -> pterm | PBind : forall (t : ptype), pprimop t -> (V t -> pterm) -> pterm with pprimop : ptype -> Type := | PVar : forall (t : ptype), V t -> pprimop t | PTrue : pprimop PBool | PFalse : pprimop PBool | PAbs : forall (t : ptype), (V t -> pterm) -> pprimop (PCont t) | PPair : forall (t1 t2 : ptype), V t1 -> V t2 -> pprimop (PProd t1 t2) | PFst : forall (t1 t2 : ptype), V (PProd t1 t2) -> pprimop t1 | PSnd : forall (t1 t2 : ptype), V (PProd t1 t2) -> pprimop t2. (* Arguments PVar {t}. *) End cpsterm. Arguments PAbs {V res t}. Arguments PPair {V res t1 t2}. Arguments PFst {V res t1 t2}. Arguments PSnd {V res t1 t2}. Check PAbs. (* Translation *) Section splices. Fixpoint splice {V : ptype -> Type} {res1 res2 : ptype} (e1: pterm) (e2: V res1 -> pterm) : @pterm V res2 := match e1 with | PHalt v => e2 v | PApp _ f x => PApp _ f x | PBind _ p f => PBind _ (splicePrim p e2) (fun x => splice (f x) e2) end with splicePrim {V : ptype -> Type} {res1 res2 t : ptype} (p : @pprimop V res1 t) (e2 : V res1 -> @pterm V res2) : @pprimop V res2 t := match p with | PVar _ v => PVar _ v | PTrue => PTrue | PFalse => PFalse (*| PAbs t f => *) | PAbs f => PAbs (fun x => splice (f x) e2) | PPair v1 v2 => PPair v1 v2 | PFst v => PFst v | PSnd v => PSnd v end. End splices. Fixpoint cpsType (t : type) : ptype := match t with | Bool => PBool | Arrow t1 t2 => PCont (PProd (cpsType t1) (PCont (cpsType t2))) end. Notation "let x := e1 in e2" := (splice e1 (fun x => e2)) (at level 80). Check pterm. Check @term. (* #+BEGIN_OUTPUT (Info) term : type -> Type where ?V : [ |- type -> Type] #+END_OUTPUT (Info) *) Check PHalt. (* #+BEGIN_OUTPUT (Info) PHalt : ?V ?res -> pterm where ?V : [ |- ptype -> Type] ?res : [ |- ptype] #+END_OUTPUT (Info) *) Section translation. Variable V : ptype -> Type. Notation V' := (fun (t : type) => V (cpsType t)). Notation "x <-- e1 ; e2" := (splice e1 (fun x => e2)) (at level 76). (* letTerm *) Notation "x <- p ; e" := (PBind p (fun x => e)) (at level 76). (* letBind *) Notation "\ x , e" := (PAbs (fun x => e)) (at level 78). (* fn/λ *) (* Notation "'letTerm' x ':=' e1 'inside' e2" := (splice e1 (fun x => e2)) (at level 70). Notation "'letBind' x ':=' e1 'inside' e2" := (PBind e1 (fun x => e2)) (at level 70). Notation "'fn' x ':=' e" := (PAbs (fun x => e)) (at level 70). *) Fixpoint cpsTerm {t : type} (e : @term V' t) : @pterm V (cpsType t) := match e with | Var _ x => PHalt x | Tru => PBind PTrue (fun x => PHalt x) | Fals => PBind PFalse (fun x => PHalt x) | App _ _ e1 e2 => f <-- (cpsTerm e1) ; x <-- (cpsTerm e2) ; k <- \r, PHalt (V:=V) x ; p <- (PPair x k) ; (PApp f p) | Abs _ _ e' => (*f <- \r , *) f <- PAbs V (fun p => x <- PFst p ; k <- PSnd p ; r <-- cpsTerm (e' x0) ; PApp k r) ; PHalt f (* Let f := PAbs V (fun p => Let x := PFst p inside Let k := PSnd p inside splice (cpsTerm (e' x)) (fun x' => PApp k r)) PHalt f *) end. End translation. (* Definition foo (l m n:nat) : nat -> nat := plus n. Check foo. Check foo 3. *) Fixpoint ptermDenote {result : ptype} (e : pterm ptypeDenote result) (k : (ptypeDenote result) -> bool) : bool := match e with | PHalt _ _ v => k v | PApp _ _ _ f x => f x (* f is the continuation function *) | PBind _ _ _ p f => ptermDenote (f (pprimopDenote p k)) k end with pprimopDenote {result t : ptype} (p : pprimop ptypeDenote result t) (k : ptypeDenote result -> bool) : ptypeDenote t := match p with | PVar _ _ t v => v | PTrue _ _ => true | PFalse _ _ => false | PAbs _ _ t f => fun x => ptermDenote (f x) k | PPair _ _ t1 t2 v1 v2 => (v1, v2) | PFst _ _ t1 t2 v => fst v | PSnd _ _ t1 t2 v => snd v end. (***************************************************************) Inductive type : Type := | Bool : type | Arrow : type -> type -> type. Section term. Variable var : type -> Type. Inductive term : type -> Type := | Var: forall (t : type), var t -> term t | App: forall (t1 t2 : type), term (Arrow t1 t2) -> term t1 -> term t2 | Abs: forall (t1 t2 : type), var t1 -> term t2 -> term (Arrow t1 t2) | Tru: term Bool | Fals: term Bool. End term. (* CPS syntax *) (* Types τ ::= bool | τ→0 | τxτ *) Inductive ctype : Type := | TCBool : ctype | TCCont : ctype -> ctype | TCUnit : ctype | TCProd : ctype -> ctype -> ctype. Section var. Variable var : ctype -> Type. Variable result : ctype. Inductive cterm : Type := (* CPS over *) | CHalt : var result -> cterm | CApp : forall (t : ctype), var (TCCont t) -> var t -> cterm (* let binding *) | CBind : forall (t : ctype), primop t -> (var t -> cterm) -> cterm with primop : ctype -> Type := | CopVar : forall (t : ctype), var t -> primop t | CopTru : primop TCBool | CopFals : primop TCBool | CopAbs : forall (t : ctype), (var t -> cterm) -> primop (TCCont t) | CopPair : forall (t1 t2 : ctype), var t1 -> var t2 -> primop (TCProd t1 t2) | CopFst : forall (t1 t2 : ctype), var (TCProd t1 t2) -> primop t1 | CopSnd : forall (t1 t2 : ctype), var (TCProd t1 t2) -> primop t2. End var. (* CPS types to coq types *) Fixpoint ctypeDenote (t : ctype) : Type := match t with | TCBool => bool (* why to bool ?*) | TCCont t' => ctypeDenote t' -> bool | TCUnit => unit | TCProd t1 t2 => ((ctypeDenote t1) * (ctypeDenote t2))%type end. Check Var. (* Var : forall (var : type -> Type) (t : type), var t -> term var t *) Check App. (* App : forall (var : type -> Type) (t1 t2 : type), term var (Arrow t1 t2) -> term var t1 -> term var t2 *) Fixpoint ctermDenote (t : cterm) : Type := match t with | CHalt : var result -> cterm | CApp : forall (t : ctype), var (TCCont t) -> var t -> cterm | CBind : forall (t : ctype), primop t -> (var t -> cterm) -> cterm end with primopDenote (result: ) (t: cterm) (op : primop) : ctermDenote t (**********************************************************) Inductive type : Type := | Nat: type | Func: type -> type -> type. (* HOAS. Strict positivity failed! Inductive term : type -> Type := | Const: term Nat | Plus: term Nat -> term Nat -> term Nat | Abs: forall (t1 t2 : type), term t1 -> term t2 -> term (Func t1 t2) | App: forall (t1 t2 : type), term (Func t1 t2) -> term t1 -> term t2 (* let x = e1 in e2 *) (* λx.e2) e1 *) (* Let e1 (λx.e2) *) | Let: forall (t1 t2: type), term t1 -> (term t1 -> term t2) -> term t2. *) Inductive term (var: type -> Type) := | Var : forall t:type, var t -> term (var t). Section var. Variable var : type -> Type. Inductive term : type -> Type := | Var : forall t:type, var t -> term t. (*************************************************************) (* Inductive term: Type := | App: term -> term -> term | Abs: (term -> term) -> term. *) (* PHOAS *) Inductive term (T: Type) : Type := | Var: T -> term T | App: term T -> term T -> term T | Abs: (T -> term T) -> term T. Require Import List. Import ListNotations. Inductive member {A : Type} (elem: A) : list A -> Type := | First : forall ls: list A, member elem (elem: ls). | Next: forall (x:A) (ls:list A), member ls -> member (x :: ls).