(** This file is a slight modification of ImpSimpl.v from Adam Chilipala's FRAP: *) From stdpp Require Import stringmap natmap. (** Here's some appropriate syntax for expressions (side-effect-free) of a simple imperative language with a mutable memory. *) Inductive exp := | Const (n : nat) | Var (x : string) | Plus (e1 e2 : exp) | Minus (e1 e2 : exp) | Mult (e1 e2 : exp). (** Those were the expressions of numeric type. Here are the Boolean expressions. *) Inductive bexp := | Equal (e1 e2 : exp) | Less (e1 e2 : exp). Definition var := string. Definition valuation := stringmap nat. Inductive cmd := | Skip | Assign (x : var) (e : exp) | Seq (c1 c2 : cmd) | If_ (be : bexp) (then_ else_ : cmd) | While_ (be : bexp) (body : cmd). (** Shorthand notation for looking up in a finite map, returning zero if the key is not found *) Notation "m $! k" := (match m !! k with Some n => n | None => O end) (at level 30). (** Start of expression semantics: meaning of expressions *) Fixpoint eval (e : exp) (v : valuation) : nat := match e with | Const n => n | Var x => v $! x | Plus e1 e2 => eval e1 v + eval e2 v | Minus e1 e2 => eval e1 v - eval e2 v | Mult e1 e2 => eval e1 v * eval e2 v end. (** Meaning of Boolean expressions *) Fixpoint beval (b : bexp) (v : valuation) : bool := match b with | Equal e1 e2 => beq_nat (eval e1 v) (eval e2 v) | Less e1 e2 => negb (eval e2 v <=? eval e1 v) end. Inductive exec : valuation -> cmd -> valuation -> Prop := | ExSkip : forall v, exec v Skip v | ExAssign : forall v x e, exec v (Assign x e) (<[x := eval e v]>v) | ExSeq : forall v1 c1 v2 c2 v3, exec v1 c1 v2 -> exec v2 c2 v3 -> exec v1 (Seq c1 c2) v3 | ExIfTrue : forall v1 b c1 c2 v2, beval b v1 = true -> exec v1 c1 v2 -> exec v1 (If_ b c1 c2) v2 | ExIfFalse : forall v1 b c1 c2 v2, beval b v1 = false -> exec v1 c2 v2 -> exec v1 (If_ b c1 c2) v2 | ExWhileFalse : forall v b c, beval b v = false -> exec v (While_ b c) v | ExWhileTrue : forall v1 b c v2 v3, beval b v1 = true -> exec v1 c v2 -> exec v2 (While_ b c) v3 -> exec v1 (While_ b c) v3. Open Scope string_scope. (* BEGIN syntax macros that won't be explained *) Coercion Const : nat >-> exp. Coercion Var : string >-> exp. Infix "+" := Plus : cmd_scope. Infix "-" := Minus : cmd_scope. Infix "*" := Mult : cmd_scope. Infix "=" := Equal : cmd_scope. Infix "<" := Less : cmd_scope. Definition set (dst src : exp) : cmd := match dst with | Var dst' => Assign dst' src | _ => Assign "Bad LHS" 0 end. Infix "<-" := set (no associativity, at level 70) : cmd_scope. Infix ";;;" := Seq (right associativity, at level 70) : cmd_scope. Notation "'when' b 'then' then_ 'else' else_ 'done'" := (If_ b then_ else_) (at level 75, b at level 0). Notation "'while' b 'loop' body 'done'" := (While_ b body) (at level 75). Delimit Scope cmd_scope with cmd. Infix "+" := plus : reset_scope. Infix "-" := minus : reset_scope. Infix "*" := mult : reset_scope. Infix "=" := eq : reset_scope. Infix "<" := lt : reset_scope. Delimit Scope reset_scope with reset. Open Scope reset_scope. (* END macros *)