playground/coq/cpdt/typed-stack-mc.v

172 lines
4.2 KiB
Coq

Inductive type : Set :=
| Nat
| Bool.
Check Nat.
Definition typeDenote (t:type) : Type :=
match t with
| Nat => nat
| Bool => bool
end.
Compute (typeDenote Nat).
Inductive binop : type -> type -> type -> Set :=
| Plus : binop Nat Nat Nat
| Mult : binop Nat Nat Nat
| Eq : forall t:type, binop t t Bool
| Lt : binop Nat Nat Bool.
Check Plus.
Definition binopDenote {t1 t2 t3 : type} (op:binop t1 t2 t3)
: (typeDenote t1) -> (typeDenote t2) -> (typeDenote t3) :=
match op with
| Plus => Nat.add
| Mult => Nat.mul
| Eq Nat => Nat.eqb
| Eq Bool =>
fun (x y:bool) => negb (xorb x y)
| Lt => Nat.ltb
end.
Compute binopDenote Plus.
Inductive exp : type -> Set :=
| NConst : nat -> exp Nat
| BConst : bool -> exp Bool
| Binop : forall (t1 t2 t3:type),
binop t1 t2 t3 -> exp t1 -> exp t2 -> exp t3.
Arguments Binop {t1 t2 t3}.
Check BConst.
Check Binop Plus (NConst 3) (NConst 4).
Fixpoint expDenote {t:type} (e:exp t) : typeDenote t :=
match e with
| NConst n => n
| BConst b => b
| Binop b e1 e2 => (binopDenote b) (expDenote e1) (expDenote e2)
end.
Compute expDenote (Binop Plus (NConst 3) (NConst 4)).
(* Target *)
Require Import List.
Import ListNotations.
Definition stack := list type.
Inductive instr : stack -> stack -> Set :=
| INConst : forall s:stack,
nat -> instr s (Nat::s)%list
| IBConst : forall s:stack,
bool -> instr s (Bool::s)%list
| IBinop : forall (t1 t2 t3:type) (b:binop t1 t2 t3) (s:stack),
instr (t1::t2::s)%list (t3::s)%list.
Arguments IBinop {t1 t2 t3}.
Check INConst [] 3.
Check IBConst [Nat] true.
(* Converts a list of [type] to an n-tuple of [Type].
end is [unit] *)
Fixpoint valstack (s:stack) : Type :=
match s with
| [] => unit
| (t::ts) => (typeDenote t) * (valstack ts)
end.
Compute valstack [Nat; Bool].
(* = (nat * (bool * unit))%type : Type *)
Definition instrDenote {s1 s2:stack} (i:instr s1 s2)
: valstack s1 -> valstack s2 :=
match i with
| INConst _ n =>
fun vs => (n, vs)
| IBConst _ b =>
fun vs => (b, vs)
| IBinop op _ =>
fun vs =>
let '(t1, (t2, ts)) := vs in
((binopDenote op) t1 t2, ts)
end.
Compute (3, (nat,unit)).
Compute (instrDenote (INConst [] 3)).
Compute instrDenote (IBConst [Nat] true).
Inductive prog : stack -> stack -> Set :=
| pnil : forall s:stack, prog s s
| pcons : forall s1 s2 s3:stack,
instr s1 s2 -> prog s2 s3 -> prog s1 s3.
Arguments pcons {s1 s2 s3}.
Check pnil [].
Check pcons (INConst [] 3) (pnil [Nat]).
Check pcons (INConst [Bool] 3) (pnil [Nat; Bool]).
Check pcons (INConst _ 3) (pnil _).
Check pcons (IBConst _ true) (pcons (INConst _ 3) (pnil _)).
(* Check pcons (INConst [] 3) (pnil []). ✗ *)
Fixpoint progDenote {s1 s2 : stack} (p:prog s1 s2)
: valstack s1 -> valstack s2 :=
match p with
| pnil _ =>
fun vs => vs
| pcons i p' =>
fun vs =>
(progDenote p' (instrDenote i vs))
end.
Check progDenote (pnil []).
Compute (valstack []).
Compute (valstack [Bool]).
Compute progDenote (pcons (INConst [] 3) (pnil [Nat])).
Compute (instrDenote (INConst [] 3)).
Compute (instrDenote (INConst [] 3) tt).
Compute (instrDenote (INConst [Bool] 3)).
Compute (instrDenote (INConst [Bool] 3) (true,tt)).
Compute (instrDenote (INConst [Nat] 3) (2, tt)).
Compute progDenote (pcons (INConst [Nat] 3) (pnil [Nat; Nat]))
(instrDenote (INConst [Nat] 3) (2, tt)).
Check progDenote.
(*
progDenote
: prog ?s1 ?s2 -> valstack ?s1 -> valstack ?s2
*)
Check instrDenote.
(*
instrDenote
: instr ?s1 ?s2 -> valstack ?s1 -> valstack ?s2
*)
(* Concat 2 programs *)
Fixpoint pconcat {s1 s2 s3:stack} (p:prog s1 s2)
: prog s2 s3 -> prog s1 s3 :=
match p with
| pnil _ => fun p' => p'
| pcons i pr => fun p' => pcons i (pconcat pr p')
end.
(*
| NConst n => pcons (INConst [] n) (pnil [Nat])
| BConst b => pcons (IBConst [] b) (pnil [Bool])
*)
Fixpoint pcompile {t:type} (e:exp t) (s:stack) : prog s (t::s) :=
match e with
| NConst n => pcons (INConst s n) (pnil (Nat::s))
| BConst b => pcons (IBConst s b) (pnil (Bool::s))
| Binop bop e1 e2 => pconcat (pcompile e2 _)
(pconcat (pcompile e1 _)
(*(pcons (IBinop bop s) (pnil _)))*)
(pcons (IBinop bop s) (pnil (_::s))))
(*
(pcons (@IBinop t1 t2 t3 bop s) (pnil (t3::s))))
*)
end.