(** * Imp: Simple Imperative Programs *) (** In this chapter, we'll take a more serious look at how to use Coq to study interesting things outside of itself. Our case study is a _simple imperative programming language_ called Imp, embodying a tiny core fragment of conventional mainstream languages such as C and Java. Here is a familiar mathematical function written in Imp. Z ::= X;; Y ::= 1;; WHILE ! (Z = 0) DO Y ::= Y * Z;; Z ::= Z - 1 END *) (** This chapter looks at how to define the _syntax_ and _semantics_ of Imp; further chapters in _Programming Language Foundations_ (_Software Foundations_, volume 2) develop a theory of _program equivalence_ and introduce _Hoare Logic_, a widely used logic for reasoning about imperative programs. *) Set Warnings "-notation-overridden,-parsing". Require Import Coq.Bool.Bool. Require Import Coq.Init.Nat. Require Import Coq.Arith.Arith. Require Import Coq.Arith.EqNat. Require Import Coq.omega.Omega. Require Import Coq.Lists.List. Import ListNotations. From LF Require Import Maps. (* ################################################################# *) (** * Arithmetic and Boolean Expressions *) (** We'll present Imp in three parts: first a core language of _arithmetic and boolean expressions_, then an extension of these expressions with _variables_, and finally a language of _commands_ including assignment, conditions, sequencing, and loops. *) (* ================================================================= *) (** ** Syntax *) Module AExp. (** _Abstract syntax trees_ for arithmetic and boolean expressions: *) Inductive aexp : Type := | ANum (n : nat) | APlus (a1 a2 : aexp) | AMinus (a1 a2 : aexp) | AMult (a1 a2 : aexp). Inductive bexp : Type := | BTrue | BFalse | BEq (a1 a2 : aexp) | BLe (a1 a2 : aexp) | BNot (b : bexp) | BAnd (b1 b2 : bexp). (** For comparison, here's a conventional BNF (Backus-Naur Form) grammar defining the same abstract syntax: a ::= nat | a + a | a - a | a * a b ::= true | false | a = a | a <= a | not b | b and b *) (* ================================================================= *) (** ** Evaluation *) (** _Evaluating_ an arithmetic expression produces a number. *) Fixpoint aeval (a : aexp) : nat := match a with | ANum n => n | APlus a1 a2 => (aeval a1) + (aeval a2) | AMinus a1 a2 => (aeval a1) - (aeval a2) | AMult a1 a2 => (aeval a1) * (aeval a2) end. Example test_aeval1: aeval (APlus (ANum 2) (ANum 2)) = 4. Proof. reflexivity. Qed. (** Similarly, evaluating a boolean expression yields a boolean. *) Fixpoint beval (b : bexp) : bool := match b with | BTrue => true | BFalse => false | BEq a1 a2 => (aeval a1) =? (aeval a2) | BLe a1 a2 => (aeval a1) <=? (aeval a2) | BNot b1 => negb (beval b1) | BAnd b1 b2 => andb (beval b1) (beval b2) end. (* QUIZ *) (** What does the following expression evaluate to? aeval (APlus (ANum 3) (AMinus (ANum 4) (ANum 1))) (1) true (2) false (3) 0 (4) 3 (5) 6 *) (* /QUIZ *) (* ================================================================= *) (** ** Optimization *) Fixpoint optimize_0plus (a:aexp) : aexp := match a with | ANum n => ANum n | APlus (ANum 0) e2 => optimize_0plus e2 | APlus e1 e2 => APlus (optimize_0plus e1) (optimize_0plus e2) | AMinus e1 e2 => AMinus (optimize_0plus e1) (optimize_0plus e2) | AMult e1 e2 => AMult (optimize_0plus e1) (optimize_0plus e2) end. Example test_optimize_0plus: optimize_0plus (APlus (ANum 2) (APlus (ANum 0) (APlus (ANum 0) (ANum 1)))) = APlus (ANum 2) (ANum 1). Proof. reflexivity. Qed. Theorem optimize_0plus_sound: forall a, aeval (optimize_0plus a) = aeval a. Proof. intros a. induction a. - (* ANum *) reflexivity. - (* APlus *) destruct a1 eqn:Ea1. + (* a1 = ANum n *) destruct n eqn:En. * (* n = 0 *) simpl. apply IHa2. * (* n <> 0 *) simpl. rewrite IHa2. reflexivity. + (* a1 = APlus a1_1 a1_2 *) simpl. simpl in IHa1. rewrite IHa1. rewrite IHa2. reflexivity. + (* a1 = AMinus a1_1 a1_2 *) simpl. simpl in IHa1. rewrite IHa1. rewrite IHa2. reflexivity. + (* a1 = AMult a1_1 a1_2 *) simpl. simpl in IHa1. rewrite IHa1. rewrite IHa2. reflexivity. - (* AMinus *) simpl. rewrite IHa1. rewrite IHa2. reflexivity. - (* AMult *) simpl. rewrite IHa1. rewrite IHa2. reflexivity. Qed. (* ################################################################# *) (** * Coq Automation *) (** That last proof was getting a little repetitive. Time to learn a few more Coq tricks... *) (* ================================================================= *) (** ** Tacticals *) (** _Tacticals_ is Coq's term for tactics that take other tactics as arguments -- "higher-order tactics," if you will. *) (* ----------------------------------------------------------------- *) (** *** The [try] Tactical *) (** If [T] is a tactic, then [try T] is a tactic that is just like [T] except that, if [T] fails, [try T] _successfully_ does nothing at all (instead of failing). *) Theorem silly1 : forall ae, aeval ae = aeval ae. Proof. try reflexivity. (* this just does [reflexivity] *) Qed. Theorem silly2 : forall (P : Prop), P -> P. Proof. intros P HP. try reflexivity. (* just [reflexivity] would have failed *) apply HP. (* we can still finish the proof in some other way *) Qed. (* ----------------------------------------------------------------- *) (** *** The [;] Tactical (Simple Form) *) (** In its most common form, the [;] tactical takes two tactics as arguments. The compound tactic [T;T'] first performs [T] and then performs [T'] on _each subgoal_ generated by [T]. *) (** For example: *) Lemma foo : forall n, 0 <=? n = true. Proof. intros. destruct n eqn:E. (* Leaves two subgoals, which are discharged identically... *) - (* n=0 *) simpl. reflexivity. - (* n=Sn' *) simpl. reflexivity. Qed. (** We can simplify this proof using the [;] tactical: *) Lemma foo' : forall n, 0 <=? n = true. Proof. intros. (* [destruct] the current goal *) destruct n; (* then [simpl] each resulting subgoal *) simpl; (* and do [reflexivity] on each resulting subgoal *) reflexivity. Qed. (** Using [try] and [;] together, we can get rid of the repetition in the proof that was bothering us a little while ago. *) Theorem optimize_0plus_sound': forall a, aeval (optimize_0plus a) = aeval a. Proof. intros a. induction a; (* Most cases follow directly by the IH... *) try (simpl; rewrite IHa1; rewrite IHa2; reflexivity). (* ... but the remaining cases -- ANum and APlus -- are different: *) - (* ANum *) reflexivity. - (* APlus *) destruct a1 eqn:Ea1; (* Again, most cases follow directly by the IH: *) try (simpl; simpl in IHa1; rewrite IHa1; rewrite IHa2; reflexivity). (* The interesting case, on which the [try...] does nothing, is when [e1 = ANum n]. In this case, we have to destruct [n] (to see whether the optimization applies) and rewrite with the induction hypothesis. *) + (* a1 = ANum n *) destruct n eqn:En; simpl; rewrite IHa2; reflexivity. Qed. (** This proof can be further improved by removing the trivial first case (for [e = ANum n]). *) Theorem optimize_0plus_sound'': forall a, aeval (optimize_0plus a) = aeval a. Proof. intros a. induction a; (* Most cases follow directly by the IH *) try (simpl; rewrite IHa1; rewrite IHa2; reflexivity); (* ... or are immediate by definition *) try reflexivity. (* The interesting case is when a = APlus a1 a2. *) - (* APlus *) destruct a1; try (simpl; simpl in IHa1; rewrite IHa1; rewrite IHa2; reflexivity). + (* a1 = ANum n *) destruct n; simpl; rewrite IHa2; reflexivity. Qed. (* ----------------------------------------------------------------- *) (** *** The [repeat] Tactical *) (** The [repeat] tactical takes another tactic and keeps applying this tactic until it fails. Here is an example showing that [8] is in a long list using repeat. *) Theorem In8 : In 8 [1;2;3;4;5;6;7;8;9;10]. Proof. repeat (try (left; reflexivity); right). Qed. (* ================================================================= *) (** ** Defining New Tactic Notations *) (** Coq also provides several ways of "programming" tactic scripts: - [Tactic Notation]: syntax extension for tactics (good for simple macros) - [Ltac]: scripting language for tactics (good for more sophisticated proof engineering) - OCaml tactic scripting API (only for wizards) An example [Tactic Notation]: *) Tactic Notation "simpl_and_try" tactic(c) := simpl; try c. (* ================================================================= *) (** ** The [omega] Tactic *) (** The [omega] tactic implements a decision procedure for a subset of first-order logic called _Presburger arithmetic_. It is based on the Omega algorithm invented by William Pugh [Pugh 1991] (in Bib.v). If the goal is a universally quantified formula made out of - numeric constants, addition ([+] and [S]), subtraction ([-] and [pred]), and multiplication by constants (this is what makes it Presburger arithmetic), - equality ([=] and [<>]) and ordering ([<=]), and - the logical connectives [/\], [\/], [~], and [->], then invoking [omega] will either solve the goal or fail, meaning that the goal is actually false. (If the goal is _not_ of this form, [omega] will also fail.) *) Example silly_presburger_example : forall m n o p, m + n <= n + o /\ o + 3 = p + 3 -> m <= p. Proof. intros. omega. Qed. (* ================================================================= *) (** ** A Few More Handy Tactics *) (** Finally, here are some miscellaneous tactics that you may find convenient. - [clear H]: Delete hypothesis [H] from the context. - [subst x]: Find an assumption [x = e] or [e = x] in the context, replace [x] with [e] throughout the context and current goal, and clear the assumption. - [subst]: Substitute away _all_ assumptions of the form [x = e] or [e = x]. - [rename... into...]: Change the name of a hypothesis in the proof context. For example, if the context includes a variable named [x], then [rename x into y] will change all occurrences of [x] to [y]. - [assumption]: Try to find a hypothesis [H] in the context that exactly matches the goal; if one is found, behave like [exact H]. - [contradiction]: Try to find a hypothesis [H] in the current context that is logically equivalent to [False]. If one is found, solve the goal. - [constructor]: Try to find a constructor [c] (from some [Inductive] definition in the current environment) that can be applied to solve the current goal. If one is found, behave like [apply c]. We'll see examples as we go along. *) (* ################################################################# *) (** * Evaluation as a Relation *) (** We have presented [aeval] and [beval] as functions defined by [Fixpoint]s. Another way to think about evaluation -- one that we will see is often more flexible -- is as a _relation_ between expressions and their values. This leads naturally to [Inductive] definitions like the following one for arithmetic expressions... *) Module aevalR_first_try. Inductive aevalR : aexp -> nat -> Prop := | E_ANum n : aevalR (ANum n) n | E_APlus (e1 e2: aexp) (n1 n2: nat) : aevalR e1 n1 -> aevalR e2 n2 -> aevalR (APlus e1 e2) (n1 + n2) | E_AMinus (e1 e2: aexp) (n1 n2: nat) : aevalR e1 n1 -> aevalR e2 n2 -> aevalR (AMinus e1 e2) (n1 - n2) | E_AMult (e1 e2: aexp) (n1 n2: nat) : aevalR e1 n1 -> aevalR e2 n2 -> aevalR (AMult e1 e2) (n1 * n2). (** A standard notation for "evaluates to": *) Notation "e '\\' n" := (aevalR e n) (at level 50, left associativity) : type_scope. End aevalR_first_try. (** If we "reserve" the notation in advance, we can even use it in the definition: *) Reserved Notation "e '\\' n" (at level 50, left associativity). Inductive aevalR : aexp -> nat -> Prop := | E_ANum n : (ANum n) \\ n | E_APlus e1 e2 n1 n2 : (e1 \\ n1) -> (e2 \\ n2) -> (APlus e1 e2) \\ (n1 + n2) | E_AMinus e1 e2 n1 n2 : (e1 \\ n1) -> (e2 \\ n2) -> (AMinus e1 e2) \\ (n1 - n2) | E_AMult e1 e2 n1 n2 : (e1 \\ n1) -> (e2 \\ n2) -> (AMult e1 e2) \\ (n1 * n2) where "e '\\' n" := (aevalR e n) : type_scope. (* ================================================================= *) (** ** Inference Rule Notation *) (** For example, the constructor [E_APlus]... | E_APlus : forall (e1 e2: aexp) (n1 n2: nat), aevalR e1 n1 -> aevalR e2 n2 -> aevalR (APlus e1 e2) (n1 + n2) ...would be written like this as an inference rule: e1 \\ n1 e2 \\ n2 -------------------- (E_APlus) APlus e1 e2 \\ n1+n2 *) (** There is nothing very deep going on here: - rule name corresponds to a constructor name - above the line are premises - below the line is conclusion - _metavariables_ like [e1] and [n1] are implicitly universally quantified - the whole collection of rules is implicitly wrapped in an [Inductive] (sometimes we write this slightly more explicitly, as "...closed under these rules...") *) End AExp. (* ================================================================= *) (** ** Computational vs. Relational Definitions *) (** We have seen that we can use either relations or functions to give a semantics expression for evaluation, with different tradeoffs. But in general, relational semantics (being lower level and more flexible) are sometimes the only reasonable option... *) Module aevalR_division. (* ----------------------------------------------------------------- *) (** *** Adding division *) (** For example, suppose that we wanted to extend the arithmetic operations by considering also a division operation:*) Inductive aexp : Type := | ANum (n : nat) | APlus (a1 a2 : aexp) | AMinus (a1 a2 : aexp) | AMult (a1 a2 : aexp) | ADiv (a1 a2 : aexp). (* <--- new *) (** Extending the definition of [aeval] to handle this new operation would not be straightforward (what should we return as the result of [ADiv (ANum 5) (ANum 0)]?). But extending [aevalR] is straightforward. *) (* ----------------------------------------------------------------- *) (** *** Adding division, relationally *) Reserved Notation "e '\\' n" (at level 50, left associativity). Inductive aevalR : aexp -> nat -> Prop := | E_ANum : forall (n:nat), (ANum n) \\ n | E_APlus : forall (a1 a2: aexp) (n1 n2 : nat), (a1 \\ n1) -> (a2 \\ n2) -> (APlus a1 a2) \\ (n1 + n2) | E_AMinus : forall (a1 a2: aexp) (n1 n2 : nat), (a1 \\ n1) -> (a2 \\ n2) -> (AMinus a1 a2) \\ (n1 - n2) | E_AMult : forall (a1 a2: aexp) (n1 n2 : nat), (a1 \\ n1) -> (a2 \\ n2) -> (AMult a1 a2) \\ (n1 * n2) | E_ADiv : forall (a1 a2: aexp) (n1 n2 n3: nat), (a1 \\ n1) -> (a2 \\ n2) -> (n2 > 0) -> (mult n2 n3 = n1) -> (ADiv a1 a2) \\ n3 where "a '\\' n" := (aevalR a n) : type_scope. End aevalR_division. Module aevalR_extended. (* ----------------------------------------------------------------- *) (** *** Adding Nondeterminism *) (** Suppose, instead, that we want to extend the arithmetic operations by a nondeterministic number generator [any] that, when evaluated, may yield any number. (Note that this is not the same as making a _probabilistic_ choice among all possible numbers -- we're not specifying any particular distribution of results, but just saying what results are _possible_.) *) Reserved Notation "e '\\' n" (at level 50, left associativity). Inductive aexp : Type := | AAny (* <--- NEW *) | ANum (n : nat) | APlus (a1 a2 : aexp) | AMinus (a1 a2 : aexp) | AMult (a1 a2 : aexp). (** Again, extending [aeval] would be tricky, since now evaluation is _not_ a deterministic function from expressions to numbers, but extending [aevalR] is no problem... *) (* ----------------------------------------------------------------- *) (** *** Adding nondeterminism, relationally *) Inductive aevalR : aexp -> nat -> Prop := | E_Any : forall (n:nat), AAny \\ n (* <--- new *) | E_ANum : forall (n:nat), (ANum n) \\ n | E_APlus : forall (a1 a2: aexp) (n1 n2 : nat), (a1 \\ n1) -> (a2 \\ n2) -> (APlus a1 a2) \\ (n1 + n2) | E_AMinus : forall (a1 a2: aexp) (n1 n2 : nat), (a1 \\ n1) -> (a2 \\ n2) -> (AMinus a1 a2) \\ (n1 - n2) | E_AMult : forall (a1 a2: aexp) (n1 n2 : nat), (a1 \\ n1) -> (a2 \\ n2) -> (AMult a1 a2) \\ (n1 * n2) where "a '\\' n" := (aevalR a n) : type_scope. End aevalR_extended. (* ----------------------------------------------------------------- *) (** *** Tradeoffs *) (** At this point you maybe wondering: which style should I use by default? The examples above show that relational definitions are fundamentally more powerful than functional ones. For situations like these, where the thing being defined is not easy to express as a function, or indeed where it is _not_ a function, there is no choice. But what about when both styles are workable? One point in favor of relational definitions is that they can be more elegant and easier to understand. Another is that Coq automatically generates nice inversion and induction principles from [Inductive] definitions. *) (** On the other hand, functional definitions can often be more convenient: - Functions are by definition deterministic and defined on all arguments; for a relation we have to show these properties explicitly if we need them. - With functions we can also take advantage of Coq's computation mechanism to simplify expressions during proofs. Furthermore, functions can be directly "extracted" to executable code in OCaml or Haskell. *) (** Ultimately, the choice often comes down to either the specifics of a particular situation or simply a question of taste. Indeed, in large Coq developments it is common to see a definition given in _both_ functional and relational styles, plus a lemma stating that the two coincide, allowing further proofs to switch from one point of view to the other at will.*) (* ################################################################# *) (** * Expressions With Variables *) (** Let's turn our attention back to defining Imp. The next thing we need to do is to enrich our arithmetic and boolean expressions with variables. To keep things simple, we'll assume that all variables are global and that they only hold numbers. *) (* ================================================================= *) (** ** States *) (** Since we'll want to look variables up to find out their current values, we'll reuse maps from the [Maps] chapter, with [string]s as the type of variables in Imp. A _machine state_ (or just _state_) represents the current values of _all_ variables at some point in the execution of a program. *) Definition state := total_map nat. (* ================================================================= *) (** ** Syntax *) (** We can add variables to the arithmetic expressions we had before by simply adding one more constructor: *) Inductive aexp : Type := | ANum (n : nat) | AId (x : string) (* <----- NEW *) | APlus (a1 a2 : aexp) | AMinus (a1 a2 : aexp) | AMult (a1 a2 : aexp). (** Defining a few variable names as notational shorthands will make examples easier to read: *) Definition W : string := "W". Definition X : string := "X". Definition Y : string := "Y". Definition Z : string := "Z". (** The definition of [bexp]s is unchanged (except that it now refers to the new [aexp]s): *) Inductive bexp : Type := | BTrue | BFalse | BEq (a1 a2 : aexp) | BLe (a1 a2 : aexp) | BNot (b : bexp) | BAnd (b1 b2 : bexp). (* ================================================================= *) (** ** Notations *) (** To make Imp programs easier to read and write, we introduce some notations and implicit coercions. You do not need to understand what these declarations do in detail to follow this chapter. Briefly, though, the [Coercion] declaration in Coq stipulates that a function (or constructor) can be implicitly used by the type system to coerce a value of the input type to a value of the output type. For instance, the coercion declaration for [AId] allows us to use plain strings when an [aexp] is expected; the string will implicitly be wrapped with [AId]. *) (** The notations below are declared in specific _notation scopes_, in order to avoid conflicts with other interpretations of the same symbols. Again, it is not necessary to understand the details. *) Coercion AId : string >-> aexp. Coercion ANum : nat >-> aexp. Definition bool_to_bexp (b: bool) : bexp := if b then BTrue else BFalse. Coercion bool_to_bexp : bool >-> bexp. Bind Scope aexp_scope with aexp. Infix "+" := APlus : aexp_scope. Infix "-" := AMinus : aexp_scope. Infix "*" := AMult : aexp_scope. Bind Scope bexp_scope with bexp. Infix "<=" := BLe : bexp_scope. Infix "=" := BEq : bexp_scope. Infix "&&" := BAnd : bexp_scope. Notation "'!' b" := (BNot b) (at level 60) : bexp_scope. (** We can now write [3 + (X * 2)] instead of [APlus 3 (AMult X 2)], and [true && !(X <= 4)] instead of [BAnd true (BNot (BLe X 4))]. *) (* ================================================================= *) (** ** Evaluation *) (** Now we need to add an [st] parameter to both evaluation functions: *) Fixpoint aeval (st : state) (a : aexp) : nat := match a with | ANum n => n | AId x => st x (* <----- NEW *) | APlus a1 a2 => (aeval st a1) + (aeval st a2) | AMinus a1 a2 => (aeval st a1) - (aeval st a2) | AMult a1 a2 => (aeval st a1) * (aeval st a2) end. Fixpoint beval (st : state) (b : bexp) : bool := match b with | BTrue => true | BFalse => false | BEq a1 a2 => (aeval st a1) =? (aeval st a2) | BLe a1 a2 => (aeval st a1) <=? (aeval st a2) | BNot b1 => negb (beval st b1) | BAnd b1 b2 => andb (beval st b1) (beval st b2) end. (** We specialize our notation for total maps to the specific case of states, i.e. using [{ --> 0 }] as empty state. *) Notation "{ a --> x }" := (t_update { --> 0 } a x) (at level 0). Notation "{ a --> x ; b --> y }" := (t_update ({ a --> x }) b y) (at level 0). Notation "{ a --> x ; b --> y ; c --> z }" := (t_update ({ a --> x ; b --> y }) c z) (at level 0). Notation "{ a --> x ; b --> y ; c --> z ; d --> t }" := (t_update ({ a --> x ; b --> y ; c --> z }) d t) (at level 0). Notation "{ a --> x ; b --> y ; c --> z ; d --> t ; e --> u }" := (t_update ({ a --> x ; b --> y ; c --> z ; d --> t }) e u) (at level 0). Notation "{ a --> x ; b --> y ; c --> z ; d --> t ; e --> u ; f --> v }" := (t_update ({ a --> x ; b --> y ; c --> z ; d --> t ; e --> u }) f v) (at level 0). (* ################################################################# *) (** * Commands *) (** Now we are ready define the syntax and behavior of Imp _commands_ (sometimes called _statements_). *) (* ================================================================= *) (** ** Syntax *) (** Informally, commands [c] are described by the following BNF grammar. (We choose this slightly awkward concrete syntax for the sake of being able to define Imp syntax using Coq's Notation mechanism. In particular, we use [IFB] to avoid conflicting with the [if] notation from the standard library.) c ::= SKIP | x ::= a | c ;; c | IFB b THEN c ELSE c FI | WHILE b DO c END *) (** For example, here's factorial in Imp: Z ::= X;; Y ::= 1;; WHILE ! (Z = 0) DO Y ::= Y * Z;; Z ::= Z - 1 END When this command terminates, the variable [Y] will contain the factorial of the initial value of [X]. *) Inductive com : Type := | CSkip | CAss (x : string) (a : aexp) | CSeq (c1 c2 : com) | CIf (b : bexp) (c1 c2 : com) | CWhile (b : bexp) (c : com). (** As for expressions, we can use a few [Notation] declarations to make reading and writing Imp programs more convenient. *) Bind Scope com_scope with com. Notation "'SKIP'" := CSkip : com_scope. Notation "x '::=' a" := (CAss x a) (at level 60) : com_scope. Notation "c1 ;; c2" := (CSeq c1 c2) (at level 80, right associativity) : com_scope. Notation "'WHILE' b 'DO' c 'END'" := (CWhile b c) (at level 80, right associativity) : com_scope. Notation "'IFB' c1 'THEN' c2 'ELSE' c3 'FI'" := (CIf c1 c2 c3) (at level 80, right associativity) : com_scope. Open Scope com_scope. (* ----------------------------------------------------------------- *) (** *** Imp Factorial in Coq *) (** For example, here is the factorial function again, written as a formal definition to Coq: *) Definition fact_in_coq : com := Z ::= X;; Y ::= 1;; WHILE ! (Z = 0) DO Y ::= Y * Z;; Z ::= Z - 1 END. (* ================================================================= *) (** ** More Examples *) (** Assignment: *) Definition plus2 : com := X ::= X + 2. Definition XtimesYinZ : com := Z ::= X * Y. Definition subtract_slowly_body : com := Z ::= Z - 1 ;; X ::= X - 1. (* ----------------------------------------------------------------- *) (** *** Loops *) Definition subtract_slowly : com := WHILE ! (X = 0) DO subtract_slowly_body END. Definition subtract_3_from_5_slowly : com := X ::= 3 ;; Z ::= 5 ;; subtract_slowly. (* ----------------------------------------------------------------- *) (** *** An infinite loop: *) Definition loop : com := WHILE true DO SKIP END. (* ################################################################# *) (** * Evaluating Commands *) (** Next we need to define what it means to evaluate an Imp command. The fact that [WHILE] loops don't necessarily terminate makes defining an evaluation function tricky... *) (* ================================================================= *) (** ** Evaluation as a Function (Failed Attempt) *) (** Here's an attempt at defining an evaluation function for commands, omitting the [WHILE] case. *) Fixpoint ceval_fun_no_while (st : state) (c : com) : state := match c with | SKIP => st | x ::= a1 => st & { x --> (aeval st a1) } | c1 ;; c2 => let st' := ceval_fun_no_while st c1 in ceval_fun_no_while st' c2 | IFB b THEN c1 ELSE c2 FI => if (beval st b) then ceval_fun_no_while st c1 else ceval_fun_no_while st c2 | WHILE b DO c END => st (* bogus *) end. (* ================================================================= *) (** ** Evaluation as a Relation *) (** Here's a better way: define [ceval] as a _relation_ rather than a _function_ -- i.e., define it in [Prop] instead of [Type], as we did for [aevalR] above. *) (** We'll use the notation [c / st \\ st'] for the [ceval] relation: [c / st \\ st'] means that executing program [c] in a starting state [st] results in an ending state [st']. This can be pronounced "[c] takes state [st] to [st']". *) (* ----------------------------------------------------------------- *) (** *** Operational Semantics *) (** Here is an informal definition of evaluation, presented as inference rules for readability: ---------------- (E_Skip) SKIP / st \\ st aeval st a1 = n -------------------------------- (E_Ass) x := a1 / st \\ st & { x --> n } c1 / st \\ st' c2 / st' \\ st'' ------------------- (E_Seq) c1;;c2 / st \\ st'' beval st b1 = true c1 / st \\ st' ------------------------------------- (E_IfTrue) IF b1 THEN c1 ELSE c2 FI / st \\ st' beval st b1 = false c2 / st \\ st' ------------------------------------- (E_IfFalse) IF b1 THEN c1 ELSE c2 FI / st \\ st' beval st b = false ------------------------------ (E_WhileFalse) WHILE b DO c END / st \\ st beval st b = true c / st \\ st' WHILE b DO c END / st' \\ st'' --------------------------------- (E_WhileTrue) WHILE b DO c END / st \\ st'' *) Reserved Notation "c1 '/' st '\\' st'" (at level 40, st at level 39). Inductive ceval : com -> state -> state -> Prop := | E_Skip : forall st, SKIP / st \\ st | E_Ass : forall st a1 n x, aeval st a1 = n -> (x ::= a1) / st \\ st & { x --> n } | E_Seq : forall c1 c2 st st' st'', c1 / st \\ st' -> c2 / st' \\ st'' -> (c1 ;; c2) / st \\ st'' | E_IfTrue : forall st st' b c1 c2, beval st b = true -> c1 / st \\ st' -> (IFB b THEN c1 ELSE c2 FI) / st \\ st' | E_IfFalse : forall st st' b c1 c2, beval st b = false -> c2 / st \\ st' -> (IFB b THEN c1 ELSE c2 FI) / st \\ st' | E_WhileFalse : forall b st c, beval st b = false -> (WHILE b DO c END) / st \\ st | E_WhileTrue : forall st st' st'' b c, beval st b = true -> c / st \\ st' -> (WHILE b DO c END) / st' \\ st'' -> (WHILE b DO c END) / st \\ st'' where "c1 '/' st '\\' st'" := (ceval c1 st st'). (** The cost of defining evaluation as a relation instead of a function is that we now need to construct _proofs_ that some program evaluates to some result state, rather than just letting Coq's computation mechanism do it for us. *) Example ceval_example1: (X ::= 2;; IFB X <= 1 THEN Y ::= 3 ELSE Z ::= 4 FI) / { --> 0 } \\ { X --> 2 ; Z --> 4 }. Proof. (* WORK IN CLASS *) Admitted. (* QUIZ *) (** Is the following proposition provable? forall c st st', (SKIP ;; c) / st \\ st' -> c / st \\ st' (1) Yes (2) No (3) Not sure *) (* /QUIZ *) (* QUIZ *) (** Is the following proposition provable? forall c1 c2 st st', (c1;;c2) / st \\ st' -> c1 / st \\ st -> c2 / st \\ st' (1) Yes (2) No (3) Not sure *) (* /QUIZ *) (* QUIZ *) (** Is the following proposition provable? forall b c st st', (IFB b THEN c ELSE c FI) / st \\ st' -> c / st \\ st' (1) Yes (2) No (3) Not sure *) (* /QUIZ *) (* QUIZ *) (** Is the following proposition provable? forall b, (forall st, beval st b = true) -> forall c st, ~(exists st', (WHILE b DO c END) / st \\ st') (1) Yes (2) No (3) Not sure *) (* /QUIZ *) (* QUIZ *) (** Is the following proposition provable? forall b c st, ~(exists st', (WHILE b DO c END) / st \\ st') -> forall st'', beval st'' b = true (1) Yes (2) No (3) Not sure *) (* /QUIZ *) (* ================================================================= *) (** ** Determinism of Evaluation (Optional) *) Theorem ceval_deterministic: forall c st st1 st2, c / st \\ st1 -> c / st \\ st2 -> st1 = st2. Proof. intros c st st1 st2 E1 E2. generalize dependent st2. induction E1; intros st2 E2; inversion E2; subst. - (* E_Skip *) reflexivity. - (* E_Ass *) reflexivity. - (* E_Seq *) assert (st' = st'0) as EQ1. { (* Proof of assertion *) apply IHE1_1; assumption. } subst st'0. apply IHE1_2. assumption. - (* E_IfTrue, b1 evaluates to true *) apply IHE1. assumption. - (* E_IfTrue, b1 evaluates to false (contradiction) *) rewrite H in H5. discriminate H5. - (* E_IfFalse, b1 evaluates to true (contradiction) *) rewrite H in H5. discriminate H5. - (* E_IfFalse, b1 evaluates to false *) apply IHE1. assumption. - (* E_WhileFalse, b1 evaluates to false *) reflexivity. - (* E_WhileFalse, b1 evaluates to true (contradiction) *) rewrite H in H2. discriminate H2. - (* E_WhileTrue, b1 evaluates to false (contradiction) *) rewrite H in H4. discriminate H4. - (* E_WhileTrue, b1 evaluates to true *) assert (st' = st'0) as EQ1. { (* Proof of assertion *) apply IHE1_1; assumption. } subst st'0. apply IHE1_2. assumption. Qed. (* ################################################################# *) (** * Additional Exercises *) (** **** Exercise: 3 stars, optional (stack_compiler) *) (** Old HP Calculators, programming languages like Forth and Postscript, and abstract machines like the Java Virtual Machine all evaluate arithmetic expressions using a _stack_. For instance, the expression (2*3)+(3*(4-2)) would be written as 2 3 * 3 4 2 - * + and evaluated like this (where we show the program being evaluated on the right and the contents of the stack on the left): [ ] | 2 3 * 3 4 2 - * + [2] | 3 * 3 4 2 - * + [3, 2] | * 3 4 2 - * + [6] | 3 4 2 - * + [3, 6] | 4 2 - * + [4, 3, 6] | 2 - * + [2, 4, 3, 6] | - * + [2, 3, 6] | * + [6, 6] | + [12] | The goal of this exercise is to write a small compiler that translates [aexp]s into stack machine instructions. The instruction set for our stack language will consist of the following instructions: - [SPush n]: Push the number [n] on the stack. - [SLoad x]: Load the identifier [x] from the store and push it on the stack - [SPlus]: Pop the two top numbers from the stack, add them, and push the result onto the stack. - [SMinus]: Similar, but subtract. - [SMult]: Similar, but multiply. *) Inductive sinstr : Type := | SPush (n : nat) | SLoad (x : string) | SPlus | SMinus | SMult. (** Write a function to evaluate programs in the stack language. It should take as input a state, a stack represented as a list of numbers (top stack item is the head of the list), and a program represented as a list of instructions, and it should return the stack after executing the program. Test your function on the examples below. Note that the specification leaves unspecified what to do when encountering an [SPlus], [SMinus], or [SMult] instruction if the stack contains less than two elements. In a sense, it is immaterial what we do, since our compiler will never emit such a malformed program. *) Fixpoint s_execute (st : state) (stack : list nat) (prog : list sinstr) : list nat (* REPLACE THIS LINE WITH ":= _your_definition_ ." *). Admitted. Example s_execute1 : s_execute { --> 0 } [] [SPush 5; SPush 3; SPush 1; SMinus] = [2; 5]. (* FILL IN HERE *) Admitted. Example s_execute2 : s_execute { X --> 3 } [3;4] [SPush 4; SLoad X; SMult; SPlus] = [15; 4]. (* FILL IN HERE *) Admitted. (** Next, write a function that compiles an [aexp] into a stack machine program. The effect of running the program should be the same as pushing the value of the expression on the stack. *) Fixpoint s_compile (e : aexp) : list sinstr (* REPLACE THIS LINE WITH ":= _your_definition_ ." *). Admitted. (** After you've defined [s_compile], prove the following to test that it works. *) Example s_compile1 : s_compile (X - (2 * Y)) = [SLoad X; SPush 2; SLoad Y; SMult; SMinus]. (* FILL IN HERE *) Admitted. (** [] *) (** **** Exercise: 4 stars, advanced (stack_compiler_correct) *) (** Now we'll prove the correctness of the compiler implemented in the previous exercise. Remember that the specification left unspecified what to do when encountering an [SPlus], [SMinus], or [SMult] instruction if the stack contains less than two elements. (In order to make your correctness proof easier you might find it helpful to go back and change your implementation!) Prove the following theorem. You will need to start by stating a more general lemma to get a usable induction hypothesis; the main theorem will then be a simple corollary of this lemma. *) Theorem s_compile_correct : forall (st : state) (e : aexp), s_execute st [] (s_compile e) = [ aeval st e ]. Proof. (* FILL IN HERE *) Admitted. (** [] *)