(** Formalization of the Abstract Syntax Tree of a subset used in the Daisy framework **) Require Import Coq.Reals.Reals Coq.QArith.QArith. Require Import Daisy.Expressions. Require Export Daisy.Infra.ExpressionAbbrevs Daisy.Infra.NatSet. (** Next define what a program is. Currently no loops, or conditionals. Only assignments and return statement **) Inductive cmd (V:Type) :Type := Let: mType -> nat -> exp V -> cmd V -> cmd V | Ret: exp V -> cmd V. Fixpoint getRetExp (V:Type) (f:cmd V) := match f with |Let m x e g => getRetExp g | Ret e => e end. Fixpoint toRCmd (f:cmd Q) := match f with |Let m x e g => Let m x (toRExp e) (toRCmd g) |Ret e => Ret (toRExp e) end. Fixpoint toREvalCmd (f:cmd R) := match f with |Let m x e g => Let M0 x (toREval e) (toREvalCmd g) |Ret e => Ret (toREval e) end. (* UNUSED! Small Step semantics for Daisy language Inductive sstep : cmd R -> env -> R -> cmd R -> env -> Prop := let_s x e s E v eps: eval_exp eps E e v -> sstep (Let x e s) E eps s (updEnv x v E) |ret_s e E v eps: eval_exp eps E e v -> sstep (Ret e) E eps (Nop R) (updEnv 0 v E). *) (** Define big step semantics for the Daisy language, terminating on a "returned" result value **) Inductive bstep : cmd R -> env -> (nat -> option mType) -> R -> mType -> Prop := let_b m m' x e s E v res defVars: eval_exp E defVars e v m -> defVars x = Some m -> bstep s (updEnv x v E) defVars res m' -> bstep (Let m x e s) E defVars res m' |ret_b m e E v defVars: eval_exp E defVars e v m -> bstep (Ret e) E defVars v m. (** The free variables of a command are all used variables of expressions without the let bound variables **) Fixpoint freeVars V (f:cmd V) :NatSet.t := match f with | Let _ x e1 g => NatSet.remove x (NatSet.union (Expressions.usedVars e1) (freeVars g)) | Ret e => Expressions.usedVars e end. (** The defined variables of a command are all let bound variables **) Fixpoint definedVars V (f:cmd V) :NatSet.t := match f with | Let _ x _ g => NatSet.add x (definedVars g) | Ret _ => NatSet.empty end. (** The live variables of a command are all variables which occur on the right hand side of an assignment or at a return statement **) Fixpoint liveVars V (f:cmd V) :NatSet.t := match f with | Let _ _ e g => NatSet.union (usedVars e) (liveVars g) | Ret e => usedVars e end.