(** * Hoare: Hoare Logic, Part I *) Set Warnings "-notation-overridden,-parsing,-deprecated-hint-without-locality". From PLF Require Import Maps. From Coq Require Import Bool.Bool. From Coq Require Import Arith.Arith. From Coq Require Import Arith.EqNat. From Coq Require Import Arith.PeanoNat. Import Nat. From Coq Require Import Lia. From PLF Require Export Imp. (** Our goal in this chapter is to carry out some simple examples of _program verification_ -- i.e., to use the precise definition of Imp to prove formally that particular programs satisfy particular specifications of their behavior. We'll develop a reasoning system called _Floyd-Hoare Logic_ -- often shortened to just _Hoare Logic_ -- in which each of the syntactic constructs of Imp is equipped with a generic "proof rule" that can be used to reason compositionally about the correctness of programs involving this construct. *) (** Hoare Logic combines two beautiful ideas: a natural way of writing down _specifications_ of programs, and a _structured proof technique_ for proving that programs are correct with respect to such specifications -- where by "structured" we mean that the structure of proofs directly mirrors the structure of the programs that they are about. *) (* ################################################################# *) (** * Assertions *) (** An _assertion_ is a logical claim about the state of a program's memory -- formally, a property of [state]s. *) Definition Assertion := state -> Prop. (** For example, - [fun st => st X = 3] holds for states [st] in which value of [X] is [3], - [fun st => True] hold for all states, and - [fun st => False] holds for no states. *) (* QUIZ Paraphrase the following assertions in English (i.e., say which states satisfy them) (1) [fun st => st X <= st Y] (2) [fun st => st X = 3 \/ st X <= st Y] (3) [fun st => st Z * st Z <= st X /\ ~ (((S (st Z)) * (S (st Z))) <= st X)] *) (** We'll use Coq's notation features to make assertions look as much like informal math as possible. For example, instead of writing fun st => st X = m we'll usually write just X = m *) (** Given two assertions [P] and [Q], we say that [P] _implies_ [Q], written [P ->> Q], if, whenever [P] holds in some state [st], [Q] also holds. *) Definition assert_implies (P Q : Assertion) : Prop := forall st, P st -> Q st. Declare Scope hoare_spec_scope. Notation "P ->> Q" := (assert_implies P Q) (at level 80) : hoare_spec_scope. Open Scope hoare_spec_scope. (** We'll also want the "iff" variant of implication between assertions: *) Notation "P <<->> Q" := (P ->> Q /\ Q ->> P) (at level 80) : hoare_spec_scope. (* ================================================================= *) (** ** Notations for Assertions *) (** The convention described above can be implemented in Coq with a little syntax magic, using coercions and annotation scopes, much as we did with [%imp] in [Imp], to automatically lift [aexp]s, numbers, and [Prop]s into [Assertion]s when they appear in the [%assertion] scope or when Coq knows that the type of an expression is [Assertion]. There is no need to understand the details of how these notation hacks work. (We barely understand some of it ourselves!) *) Definition Aexp : Type := state -> nat. Definition assert_of_Prop (P : Prop) : Assertion := fun _ => P. Definition Aexp_of_nat (n : nat) : Aexp := fun _ => n. Definition Aexp_of_aexp (a : aexp) : Aexp := fun st => aeval st a. Coercion assert_of_Prop : Sortclass >-> Assertion. Coercion Aexp_of_nat : nat >-> Aexp. Coercion Aexp_of_aexp : aexp >-> Aexp. Add Printing Coercion Aexp_of_nat Aexp_of_aexp assert_of_Prop. Arguments assert_of_Prop /. Arguments Aexp_of_nat /. Arguments Aexp_of_aexp /. Add Printing Coercion Aexp_of_nat Aexp_of_aexp assert_of_Prop. Declare Scope assertion_scope. Bind Scope assertion_scope with Assertion. Bind Scope assertion_scope with Aexp. Delimit Scope assertion_scope with assertion. Notation assert P := (P%assertion : Assertion). Notation mkAexp a := (a%assertion : Aexp). Notation "~ P" := (fun st => ~ assert P st) : assertion_scope. Notation "P /\ Q" := (fun st => assert P st /\ assert Q st) : assertion_scope. Notation "P \/ Q" := (fun st => assert P st \/ assert Q st) : assertion_scope. Notation "P -> Q" := (fun st => assert P st -> assert Q st) : assertion_scope. Notation "P <-> Q" := (fun st => assert P st <-> assert Q st) : assertion_scope. Notation "a = b" := (fun st => mkAexp a st = mkAexp b st) : assertion_scope. Notation "a <> b" := (fun st => mkAexp a st <> mkAexp b st) : assertion_scope. Notation "a <= b" := (fun st => mkAexp a st <= mkAexp b st) : assertion_scope. Notation "a < b" := (fun st => mkAexp a st < mkAexp b st) : assertion_scope. Notation "a >= b" := (fun st => mkAexp a st >= mkAexp b st) : assertion_scope. Notation "a > b" := (fun st => mkAexp a st > mkAexp b st) : assertion_scope. Notation "a + b" := (fun st => mkAexp a st + mkAexp b st) : assertion_scope. Notation "a - b" := (fun st => mkAexp a st - mkAexp b st) : assertion_scope. Notation "a * b" := (fun st => mkAexp a st * mkAexp b st) : assertion_scope. (** Lift functions to operate on assertion expressions. *) Definition ap {X} (f : nat -> X) (x : Aexp) := fun st => f (x st). Definition ap2 {X} (f : nat -> nat -> X) (x : Aexp) (y : Aexp) (st : state) := f (x st) (y st). Module ExamplePrettyAssertions. Definition ex1 : Assertion := X = 3. Definition ex2 : Assertion := True. Definition ex3 : Assertion := False. Definition assn1 : Assertion := X <= Y. Definition assn2 : Assertion := X = 3 \/ X <= Y. Definition assn3 : Assertion := Z = ap2 max X Y. Definition assn4 : Assertion := Z * Z <= X /\ ~ (((ap S Z) * (ap S Z)) <= X). End ExamplePrettyAssertions. (* ################################################################# *) (** * Hoare Triples, Informally *) (** A _Hoare triple_ is a claim about the state before and after executing a command: {{P}} c {{Q}} This means: - If command [c] begins execution in a state satisfying assertion [P], - and if [c] eventually terminates in some final state, - then that final state will satisfy the assertion [Q]. Assertion [P] is called the _precondition_ of the triple, and [Q] is the _postcondition_. *) (** For example, - [{{X = 0}} X := X + 1 {{X = 1}}] is a valid Hoare triple, stating that command [X := X + 1] would transform a state in which [X = 0] to a state in which [X = 1]. - [forall m, {{X = m}} X := X + 1 {{X = m + 1}}], is a _proposition_ stating that the Hoare triple [{{X = m}} X := X + m {{X = m + 1}}] is valid for any choice of [m]. Note that [m] in the two assertions and the command in the middle is a reference to the _Coq_ variable [m], which is bound outside the Hoare triple. *) (* QUIZ Paraphrase the following in English. 1) {{True}} c {{X = 5}} 2) forall m, {{X = m}} c {{X = m + 5)}} 3) {{X <= Y}} c {{Y <= X}} 4) {{True}} c {{False}} 5) forall m, {{X = m}} c {{Y = real_fact m}} 6) forall m, {{X = m}} c {{(Z * Z) <= m /\ ~ (((S Z) * (S Z)) <= m)}} *) (* QUIZ Is the following Hoare triple _valid_ -- i.e., is the claimed relation between [P], [c], and [Q] true? {{True}} X := 5 {{X = 5}} (1) Yes (2) No *) (* QUIZ What about this one? {{X = 2}} X := X + 1 {{X = 3}} (1) Yes (2) No *) (* QUIZ What about this one? {{True}} X := 5; Y := 0 {{X = 5}} (1) Yes (2) No *) (* QUIZ What about this one? {{X = 2 /\ X = 3}} X := 5 {{X = 0}} (1) Yes (2) No *) (* QUIZ What about this one? {{True}} skip {{False}} (1) Yes (2) No *) (* QUIZ What about this one? {{False}} skip {{True}} (1) Yes (2) No *) (* QUIZ What about this one? {{True}} while true do skip end {{False}} (1) Yes (2) No *) (* QUIZ This one? {{X = 0}} while X = 0 do X := X + 1 end {{X = 1}} (1) Yes (2) No *) (* QUIZ This one? {{X = 1}} while X <> 0 do X := X + 1 end {{X = 100}} (1) Yes (2) No *) (* ################################################################# *) (** * Hoare Triples, Formally *) (** We can formalize valid Hoare triples in Coq as follows: *) Definition hoare_triple (P : Assertion) (c : com) (Q : Assertion) : Prop := forall st st', st =[ c ]=> st' -> P st -> Q st'. Notation "{{ P }} c {{ Q }}" := (hoare_triple P c Q) (at level 90, c custom com at level 99) : hoare_spec_scope. Check ({{True}} X := 0 {{True}}). (** **** Exercise: 1 star, standard (hoare_post_true) *) (** Prove that if [Q] holds in every state, then any triple with [Q] as its postcondition is valid. *) Theorem hoare_post_true : forall (P Q : Assertion) c, (forall st, Q st) -> {{P}} c {{Q}}. Proof. (* FILL IN HERE *) Admitted. (** [] *) (** **** Exercise: 1 star, standard (hoare_pre_false) *) (** Prove that if [P] holds in no state, then any triple with [P] as its precondition is valid. *) Theorem hoare_pre_false : forall (P Q : Assertion) c, (forall st, ~ (P st)) -> {{P}} c {{Q}}. Proof. (* FILL IN HERE *) Admitted. (** [] *) (* ################################################################# *) (** * Proof Rules *) (** Here's our plan: - introduce one "proof rule" for each Imp syntactic form - plus a couple of "structural rules" that help glue proofs together - prove these rules correct in terms of the definition of [hoare_triple] - prove programs correct using these proof rules, without ever unfolding the definition of [hoare_triple] *) (* ================================================================= *) (** ** Skip *) (** Since [skip] doesn't change the state, it preserves any assertion [P]: -------------------- (hoare_skip) {{ P }} skip {{ P }} *) Theorem hoare_skip : forall P, {{P}} skip {{P}}. Proof. intros P st st' H HP. inversion H; subst. assumption. Qed. (* ================================================================= *) (** ** Sequencing *) (** If command [c1] takes any state where [P] holds to a state where [Q] holds, and if [c2] takes any state where [Q] holds to one where [R] holds, then doing [c1] followed by [c2] will take any state where [P] holds to one where [R] holds: {{ P }} c1 {{ Q }} {{ Q }} c2 {{ R }} ---------------------- (hoare_seq) {{ P }} c1;c2 {{ R }} *) Theorem hoare_seq : forall P Q R c1 c2, {{Q}} c2 {{R}} -> {{P}} c1 {{Q}} -> {{P}} c1; c2 {{R}}. Proof. unfold hoare_triple. intros P Q R c1 c2 H1 H2 st st' H12 Pre. inversion H12; subst. eauto. Qed. (* ================================================================= *) (** ** Assignment *) (** How can we complete this triple? {{ ??? }} X := Y {{ X = 1 }} One natural possibility is: {{ Y = 1 }} X := Y {{ X = 1 }} The precondition is just the postcondition, but with [X] replaced by [Y]. *) (** How about this one? {{ ??? }} X := X + Y {{ X = 1 }} Replace [X] with [X + Y]: {{ X + Y = 1 }} X := X + Y {{ X = 1 }} This works because "equals 1" holding of [X] is guaranteed by the property "equals 1" holding of whatever is being assigned to [X]. *) (** In general, the postcondition could be some arbitrary assertion [Q], and the right-hand side of the assignment could be some arbitrary arithmetic expression [a]: {{ ??? }} X := a {{ Q }} The precondition would then be [Q], but with any occurrences of [X] in it replaced by [a]. Let's introduce a notation for this idea of replacing occurrences: Define [Q [X |-> a]] to mean "[Q] where [a] is substituted in place of [X]". This yields the Hoare logic rule for assignment: {{ Q [X |-> a] }} X := a {{ Q }} One way of reading this rule is: If you want statement [X := a] to terminate in a state that satisfies assertion [Q], then it suffices to start in a state that also satisfies [Q], except where [a] is substituted for every occurrence of [X]. *) (** Here are some valid instances of the assignment rule: {{ (X <= 5) [X |-> X + 1] }} (that is, X + 1 <= 5) X := X + 1 {{ X <= 5 }} {{ (X = 3) [X |-> 3] }} (that is, 3 = 3) X := 3 {{ X = 3 }} {{ (0 <= X /\ X <= 5) [X |-> 3] (that is, 0 <= 3 /\ 3 <= 5) X := 3 {{ 0 <= X /\ X <= 5 }} *) (** To formalize the rule, we must first formalize the idea of "substituting an expression for an Imp variable in an assertion", which we refer to as assertion substitution, or [assn_sub]. That is, intuitively, given a proposition [P], a variable [X], and an arithmetic expression [a], we want to derive another proposition [P'] that is just the same as [P] except that [P'] should mention [a] wherever [P] mentions [X]. *) (** This operation is related to the idea of substituting Imp expressions for Imp variables that we saw in [Equiv] ([subst_aexp] and friends). The difference is that, here, [P] is an arbitrary Coq assertion, so we can't directly "edit" its text. *) (** However, we can achieve the same effect by evaluating [P] in an updated state: *) Definition assn_sub X a (P:Assertion) : Assertion := fun (st : state) => P (X !-> aeval st a ; st). Notation "P [ X |-> a ]" := (assn_sub X a P) (at level 10, X at next level, a custom com). (** That is, [P [X |-> a]] stands for an assertion -- let's call it [P'] -- that is just like [P] except that, wherever [P] looks up the variable [X] in the current state, [P'] instead uses the value of the expression [a]. *) (** Now, using the substitution operation we've just defined, we can give the precise proof rule for assignment: ---------------------------- (hoare_asgn) {{Q [X |-> a]}} X := a {{Q}} *) (** We can prove formally that this rule is indeed valid. *) Theorem hoare_asgn : forall Q X a, {{Q [X |-> a]}} X := a {{Q}}. Proof. unfold hoare_triple. intros Q X a st st' HE HQ. inversion HE. subst. unfold assn_sub in HQ. assumption. Qed. (** Here's a first formal proof using this rule. *) Example assn_sub_example : {{(X < 5) [X |-> X + 1]}} X := X + 1 {{X < 5}}. Proof. apply hoare_asgn. Qed. (** Of course, what we'd probably prefer is to prove this simpler triple: {{X < 4}} X := X + 1 {{X < 5}} We will see how to do so in the next section. *) (* ================================================================= *) (** ** Consequence *) (** Sometimes the preconditions and postconditions we get from the Hoare rules won't quite be the ones we want in the particular situation at hand -- they may be logically equivalent but have a different syntactic form that fails to unify with the goal we are trying to prove, or they actually may be logically weaker (for preconditions) or stronger (for postconditions) than what we need. *) (** For instance, {{(X = 3) [X |-> 3]}} X := 3 {{X = 3}}, follows directly from the assignment rule, but {{True}} X := 3 {{X = 3}} does not. This triple is valid, but it is not an instance of [hoare_asgn] because [True] and [(X = 3) [X |-> 3]] are not syntactically equal assertions. However, they are logically _equivalent_, so if one triple is valid, then the other must certainly be as well. We can capture this observation with the following rule: {{P'}} c {{Q}} P <<->> P' ----------------------------- (hoare_consequence_pre_equiv) {{P}} c {{Q}} *) (** Taking this line of thought a bit further, we can see that strengthening the precondition or weakening the postcondition of a valid triple always produces another valid triple. This observation is captured by two _Rules of Consequence_. {{P'}} c {{Q}} P ->> P' ----------------------------- (hoare_consequence_pre) {{P}} c {{Q}} {{P}} c {{Q'}} Q' ->> Q ----------------------------- (hoare_consequence_post) {{P}} c {{Q}} *) (** Here are the formal versions: *) Theorem hoare_consequence_pre : forall (P P' Q : Assertion) c, {{P'}} c {{Q}} -> P ->> P' -> {{P}} c {{Q}}. Proof. unfold hoare_triple, "->>". intros P P' Q c Hhoare Himp st st' Heval Hpre. apply Hhoare with (st := st). - assumption. - apply Himp. assumption. Qed. Theorem hoare_consequence_post : forall (P Q Q' : Assertion) c, {{P}} c {{Q'}} -> Q' ->> Q -> {{P}} c {{Q}}. Proof. unfold hoare_triple, "->>". intros P Q Q' c Hhoare Himp st st' Heval Hpre. apply Himp. apply Hhoare with (st := st). - assumption. - assumption. Qed. (** For example, we can use the first consequence rule like this: {{ True }} ->> {{ (X = 1) [X |-> 1] }} X := 1 {{ X = 1 }} Or, formally... *) Example hoare_asgn_example1 : {{True}} X := 1 {{X = 1}}. Proof. (* WORK IN CLASS *) Admitted. (** We can also use it to prove the example mentioned earlier. {{ X < 4 }} ->> {{ (X < 5)[X |-> X + 1] }} X := X + 1 {{ X < 5 }} Or, formally ... *) Example assn_sub_example2 : {{X < 4}} X := X + 1 {{X < 5}}. Proof. (* WORK IN CLASS *) Admitted. (** Finally, here is a combined rule of consequence that allows us to vary both the precondition and the postcondition. {{P'}} c {{Q'}} P ->> P' Q' ->> Q ----------------------------- (hoare_consequence) {{P}} c {{Q}} *) Theorem hoare_consequence : forall (P P' Q Q' : Assertion) c, {{P'}} c {{Q'}} -> P ->> P' -> Q' ->> Q -> {{P}} c {{Q}}. Proof. intros P P' Q Q' c Htriple Hpre Hpost. apply hoare_consequence_pre with (P' := P'). - apply hoare_consequence_post with (Q' := Q'). + assumption. + assumption. - assumption. Qed. (* ================================================================= *) (** ** Automation *) (** Many of the proofs we have done so far with Hoare triples can be streamlined using the automation techniques that we introduced in the [Auto] chapter of _Logical Foundations_. Recall that the [auto] tactic can be told to [unfold] definitions as part of its proof search. Let's give that hint for the definitions and coercions we're using: *) Hint Unfold assert_implies hoare_triple assn_sub t_update : core. Hint Unfold assert_of_Prop Aexp_of_nat Aexp_of_aexp : core. (** Also recall that [auto] will search for a proof involving [intros] and [apply]. By default, the theorems that it will apply include any of the local hypotheses, as well as theorems in a core database. *) (** Here's a good candidate for automation: *) Theorem hoare_consequence_pre' : forall (P P' Q : Assertion) c, {{P'}} c {{Q}} -> P ->> P' -> {{P}} c {{Q}}. Proof. unfold hoare_triple, "->>". intros P P' Q c Hhoare Himp st st' Heval Hpre. apply Hhoare with (st := st). - assumption. - apply Himp. assumption. Qed. (** Tactic [eapply] will find [st] for us. *) Theorem hoare_consequence_pre''' : forall (P P' Q : Assertion) c, {{P'}} c {{Q}} -> P ->> P' -> {{P}} c {{Q}}. Proof. unfold hoare_triple, "->>". intros P P' Q c Hhoare Himp st st' Heval Hpre. eapply Hhoare. - eassumption. - apply Himp. assumption. Qed. (** The [eauto] tactic will use [eapply] as part of its proof search. So, the entire proof can actually be done in just one line. *) Theorem hoare_consequence_pre'''' : forall (P P' Q : Assertion) c, {{P'}} c {{Q}} -> P ->> P' -> {{P}} c {{Q}}. Proof. eauto. Qed. (** ...as can the same proof for the postcondition consequence rule. *) Theorem hoare_consequence_post' : forall (P Q Q' : Assertion) c, {{P}} c {{Q'}} -> Q' ->> Q -> {{P}} c {{Q}}. Proof. eauto. Qed. (** We can also use [eapply] to streamline a proof ([hoare_asgn_example1]), that we did earlier as an example of using the consequence rule: *) Example hoare_asgn_example1' : {{True}} X := 1 {{X = 1}}. Proof. eapply hoare_consequence_pre. (* no need to state an assertion *) - apply hoare_asgn. - unfold "->>", assn_sub, t_update. intros st _. simpl. reflexivity. Qed. (** The final bullet of that proof also looks like a candidate for automation. *) Example hoare_asgn_example1'' : {{True}} X := 1 {{X = 1}}. Proof. eapply hoare_consequence_pre. - apply hoare_asgn. - auto. Qed. (** Now we have quite a nice proof script: it simply identifies the Hoare rules that need to be used and leaves the remaining low-level details up to Coq to figure out. *) (** The other example of using consequence that we did earlier, [hoare_asgn_example2], requires a little more work to automate. We can streamline the first line with [eapply], but we can't just use [auto] for the final bullet, since it needs [lia]. *) Example assn_sub_example2' : {{X < 4}} X := X + 1 {{X < 5}}. Proof. eapply hoare_consequence_pre. - apply hoare_asgn. - auto. (* no progress *) unfold "->>", assn_sub, t_update. intros st H. simpl in *. lia. Qed. (** Let's introduce our own tactic to handle both that bullet and the bullet from example 1: *) Ltac assn_auto := try auto; (* as in example 1, above *) try (unfold "->>", assn_sub, t_update; intros; simpl in *; lia). (* as in example 2 *) Example assn_sub_example2'' : {{X < 4}} X := X + 1 {{X < 5}}. Proof. eapply hoare_consequence_pre. - apply hoare_asgn. - assn_auto. Qed. Example hoare_asgn_example1''': {{True}} X := 1 {{X = 1}}. Proof. eapply hoare_consequence_pre. - apply hoare_asgn. - assn_auto. Qed. (** Again, we have quite a nice proof script. All the low-level details of proofs about assertions have been taken care of automatically. Of course, [assn_auto] isn't able to prove everything we could possibly want to know about assertions -- there's no magic here! But it's good enough so far. *) (* ================================================================= *) (** ** Sequencing + Assignment *) (** Here's an example of a program involving both sequencing and assignment. Note the use of [hoare_seq] in conjunction with [hoare_consequence_pre] and the [eapply] tactic. *) Example hoare_asgn_example3 : forall (a:aexp) (n:nat), {{a = n}} X := a; skip {{X = n}}. Proof. intros a n. eapply hoare_seq. - (* right part of seq *) apply hoare_skip. - (* left part of seq *) eapply hoare_consequence_pre. + apply hoare_asgn. + assn_auto. Qed. (** Informally, a nice way of displaying a proof using the sequencing rule is as a "decorated program" where the intermediate assertion [Q] is written between [c1] and [c2]: {{ a = n }} X := a {{ X = n }}; <--- decoration for Q skip {{ X = n }} *) (** We'll come back to the idea of decorated programs in much more detail in the next chapter. *) (* ================================================================= *) (** ** Conditionals *) (** What sort of rule do we want for reasoning about conditional commands? Certainly, if the same assertion [Q] holds after executing either of the branches, then it holds after the whole conditional. So we might be tempted to write: {{P}} c1 {{Q}} {{P}} c2 {{Q}} --------------------------------- {{P}} if b then c1 else c2 {{Q}} *) (** However, this is rather weak. For example, using this rule, we cannot show {{ True }} if X = 0 then Y := 2 else Y := X + 1 end {{ X <= Y }} since the rule tells us nothing about the state in which the assignments take place in the "then" and "else" branches. *) (** Better: {{P /\ b}} c1 {{Q}} {{P /\ ~ b}} c2 {{Q}} ------------------------------------ (hoare_if) {{P}} if b then c1 else c2 end {{Q}} *) (** To make this formal, we need a way of formally "lifting" any bexp [b] to an assertion. We'll write [bassn b] for the assertion "the boolean expression [b] evaluates to [true]." *) Definition bassn b : Assertion := fun st => (beval st b = true). Coercion bassn : bexp >-> Assertion. Arguments bassn /. (** A useful fact about [bassn]: *) Lemma bexp_eval_false : forall b st, beval st b = false -> ~ ((bassn b) st). Proof. congruence. Qed. Hint Resolve bexp_eval_false : core. (** Now we can formalize the Hoare proof rule for conditionals and prove it correct. *) Theorem hoare_if : forall P Q (b:bexp) c1 c2, {{ P /\ b }} c1 {{Q}} -> {{ P /\ ~ b}} c2 {{Q}} -> {{P}} if b then c1 else c2 end {{Q}}. (** That is (unwrapping the notations): Theorem hoare_if : forall P Q b c1 c2, {{fun st => P st /\ bassn b st}} c1 {{Q}} -> {{fun st => P st /\ ~ (bassn b st)}} c2 {{Q}} -> {{P}} if b then c1 else c2 end {{Q}}. *) Proof. intros P Q b c1 c2 HTrue HFalse st st' HE HP. inversion HE; subst; eauto. Qed. (* ----------------------------------------------------------------- *) (** *** Example *) Example if_example : {{True}} if (X = 0) then Y := 2 else Y := X + 1 end {{X <= Y}}. Proof. apply hoare_if. - (* Then *) eapply hoare_consequence_pre. + apply hoare_asgn. + assn_auto. (* no progress *) unfold "->>", assn_sub, t_update, bassn. simpl. intros st [_ H]. apply eqb_eq in H. rewrite H. lia. - (* Else *) eapply hoare_consequence_pre. + apply hoare_asgn. + assn_auto. Qed. (** As we did earlier, it would be nice to eliminate all the low-level proof script that isn't about the Hoare rules. Unfortunately, the [assn_auto] tactic we wrote wasn't quite up to the job. Looking at the proof of [if_example], we can see why. We had to unfold a definition ([bassn]) and use a theorem ([eqb_eq]) that we didn't need in earlier proofs. So, let's add those into our tactic, and clean it up a little in the process. *) Ltac assn_auto' := unfold "->>", assn_sub, t_update, bassn; intros; simpl in *; try rewrite -> eqb_eq in *; (* for equalities *) auto; try lia. (** Now the proof is quite streamlined. *) Example if_example'' : {{True}} if X = 0 then Y := 2 else Y := X + 1 end {{X <= Y}}. Proof. apply hoare_if. - eapply hoare_consequence_pre. + apply hoare_asgn. + assn_auto'. - eapply hoare_consequence_pre. + apply hoare_asgn. + assn_auto'. Qed. (** We can even shorten it a little bit more. *) Example if_example''' : {{True}} if X = 0 then Y := 2 else Y := X + 1 end {{X <= Y}}. Proof. apply hoare_if; eapply hoare_consequence_pre; try apply hoare_asgn; try assn_auto'. Qed. (** For later proofs, it will help to extend [assn_auto'] to handle inequalities, too. *) Ltac assn_auto'' := unfold "->>", assn_sub, t_update, bassn; intros; simpl in *; try rewrite -> eqb_eq in *; try rewrite -> leb_le in *; (* for inequalities *) auto; try lia. (** [] *) (* ================================================================= *) (** ** While Loops *) (** The Hoare rule for [while] loops is based on the idea of an _invariant_: an assertion whose truth is guaranteed before and after executing a command. An assertion [P] is an invariant of [c] if {{P}} c {{P}} holds. Note that in the middle of executing [c], the invariant might temporarily become false, but by the end of [c], it must be restored. *) (** The Hoare while rule combines the idea of an invariant with information about when guard [b] does or does not hold. {{P /\ b}} c {{P}} --------------------------------- (hoare_while) {{P} while b do c end {{P /\ ~b}} *) Theorem hoare_while : forall P (b:bexp) c, {{P /\ b}} c {{P}} -> {{P}} while b do c end {{P /\ ~ b}}. Proof. intros P b c Hhoare st st' Heval HP. (* We proceed by induction on [Heval], because, in the "keep looping" case, its hypotheses talk about the whole loop instead of just [c]. The [remember] is used to keep the original command in the hypotheses; otherwise, it would be lost in the [induction]. By using [inversion] we clear away all the cases except those involving [while]. *) remember <{while b do c end}> as original_command eqn:Horig. induction Heval; try (inversion Horig; subst; clear Horig); eauto. Qed. (** We call [P] a _loop invariant_ of [while b do c end] if {{P /\ b}} c {{P}} is a valid Hoare triple. This means that [P] will be true at the end of the loop body whenever the loop body executes. If [P] contradicts [b], this holds trivially since the precondition is false. For instance, [X = 0] is a loop invariant of while X = 2 do X := 1 end since we will never enter the loop. *) (* QUIZ Is the assertion Y = 0 a loop invariant of the following? while X < 100 do X := X + 1 end (1) Yes (2) No *) (* QUIZ Is the assertion X = 0 a loop invariant of the following? while X < 100 do X := X + 1 end (1) Yes (2) No *) (* QUIZ Is the assertion X < Y a loop invariant of the following? while true do X := X + 1; Y := Y + 1 end (1) Yes (2) No *) (* QUIZ Is the assertion X = Y + Z a loop invariant of the following? while Y > 10 do Y := Y - 1; Z := Z + 1 end (1) Yes (2) No *) (* QUIZ Is the assertion X > 0 a loop invariant of the following? while X = 0 do X := X - 1 end (1) Yes (2) No *) (* QUIZ Is the assertion X < 100 a loop invariant of the following? while X < 100 do X := X + 1 end (1) Yes (2) No *) (* QUIZ Is the assertion X > 10 a loop invariant of the following? while X > 10 do X := X + 1 end (1) Yes (2) No *) (** If the loop never terminates, any postcondition will work. *) Theorem always_loop_hoare : forall Q, {{True}} while true do skip end {{Q}}. Proof. intros Q. eapply hoare_consequence_post. - apply hoare_while. apply hoare_post_true. auto. - simpl. intros st [Hinv Hguard]. congruence. Qed. (* ################################################################# *) (** * Summary *) (** The rules of Hoare Logic are: --------------------------- (hoare_asgn) {{Q [X |-> a]}} X:=a {{Q}} -------------------- (hoare_skip) {{ P }} skip {{ P }} {{ P }} c1 {{ Q }} {{ Q }} c2 {{ R }} ---------------------- (hoare_seq) {{ P }} c1;c2 {{ R }} {{P /\ b}} c1 {{Q}} {{P /\ ~ b}} c2 {{Q}} ------------------------------------ (hoare_if) {{P}} if b then c1 else c2 end {{Q}} {{P /\ b}} c {{P}} ----------------------------------- (hoare_while) {{P}} while b do c end {{P /\ ~ b}} {{P'}} c {{Q'}} P ->> P' Q' ->> Q ----------------------------- (hoare_consequence) {{P}} c {{Q}} *) (** Our task in this chapter has been to _define_ the rules of Hoare logic, and prove that the definitions are sound. Having done so, we can now work _within_ Hoare logic to prove that particular programs satisfy particular Hoare triples. In the next chapter, we'll see how Hoare logic is can be used to prove that more interesting programs satisfy interesting specifications of their behavior. Crucially, we will do so without ever again [unfold]ing the definition of Hoare triples -- i.e., we will take the rules of Hoare logic as a closed world for reasoning about programs. *) (* ################################################################# *) (** * Additional Exercises *) (* ================================================================= *) (** ** Havoc *) (** In this exercise, we will derive proof rules for a [HAVOC] command, which is similar to the nondeterministic [any] expression from the the [Imp] chapter. First, we enclose this work in a separate module, and recall the syntax and big-step semantics of Himp commands. *) Module Himp. Inductive com : Type := | CSkip : com | CAsgn : string -> aexp -> com | CSeq : com -> com -> com | CIf : bexp -> com -> com -> com | CWhile : bexp -> com -> com | CHavoc : string -> com. Notation "'havoc' l" := (CHavoc l) (in custom com at level 60, l constr at level 0). Notation "'skip'" := CSkip (in custom com at level 0). Notation "x := y" := (CAsgn x y) (in custom com at level 0, x constr at level 0, y at level 85, no associativity). Notation "x ; y" := (CSeq x y) (in custom com at level 90, right associativity). Notation "'if' x 'then' y 'else' z 'end'" := (CIf x y z) (in custom com at level 89, x at level 99, y at level 99, z at level 99). Notation "'while' x 'do' y 'end'" := (CWhile x y) (in custom com at level 89, x at level 99, y at level 99). Inductive ceval : com -> state -> state -> Prop := | E_Skip : forall st, st =[ skip ]=> st | E_Asgn : forall st a1 n x, aeval st a1 = n -> st =[ x := a1 ]=> (x !-> n ; st) | E_Seq : forall c1 c2 st st' st'', st =[ c1 ]=> st' -> st' =[ c2 ]=> st'' -> st =[ c1 ; c2 ]=> st'' | E_IfTrue : forall st st' b c1 c2, beval st b = true -> st =[ c1 ]=> st' -> st =[ if b then c1 else c2 end ]=> st' | E_IfFalse : forall st st' b c1 c2, beval st b = false -> st =[ c2 ]=> st' -> st =[ if b then c1 else c2 end ]=> st' | E_WhileFalse : forall b st c, beval st b = false -> st =[ while b do c end ]=> st | E_WhileTrue : forall st st' st'' b c, beval st b = true -> st =[ c ]=> st' -> st' =[ while b do c end ]=> st'' -> st =[ while b do c end ]=> st'' | E_Havoc : forall st x n, st =[ havoc x ]=> (x !-> n ; st) where "st '=[' c ']=>' st'" := (ceval c st st'). Hint Constructors ceval : core. (** The definition of Hoare triples is exactly as before. *) Definition hoare_triple (P:Assertion) (c:com) (Q:Assertion) : Prop := forall st st', st =[ c ]=> st' -> P st -> Q st'. Hint Unfold hoare_triple : core. Notation "{{ P }} c {{ Q }}" := (hoare_triple P c Q) (at level 90, c custom com at level 99) : hoare_spec_scope. (** And the precondition consequence rule is exactly as before. *) Theorem hoare_consequence_pre : forall (P P' Q : Assertion) c, {{P'}} c {{Q}} -> P ->> P' -> {{P}} c {{Q}}. Proof. eauto. Qed. (** **** Exercise: 3 stars, standard (hoare_havoc) *) (** Complete the Hoare rule for [HAVOC] commands below by defining [havoc_pre], and prove that the resulting rule is correct. *) Definition havoc_pre (X : string) (Q : Assertion) (st : total_map nat) : Prop (* REPLACE THIS LINE WITH ":= _your_definition_ ." *). Admitted. Theorem hoare_havoc : forall (Q : Assertion) (X : string), {{ havoc_pre X Q }} havoc X {{ Q }}. Proof. (* FILL IN HERE *) Admitted. (** [] *) (** **** Exercise: 3 stars, standard (havoc_post) *) (** Complete the following proof without changing any of the provided commands. If you find that it can't be completed, your definition of [havoc_pre] is probably too strong. Find a way to relax it so that [havoc_post] can be proved. Hint: the [assn_auto] tactics we've built won't help you here. You need to proceed manually. *) Theorem havoc_post : forall (P : Assertion) (X : string), {{ P }} havoc X {{ fun st => exists (n:nat), P [X |-> n] st }}. Proof. intros P X. eapply hoare_consequence_pre. - apply hoare_havoc. - (* FILL IN HERE *) Admitted. (** [] *) End Himp. (* ================================================================= *) (** ** Assert and Assume *) (** **** Exercise: 4 stars, standard (assert_vs_assume) In this exercise, we will extend IMP with two commands, [assert] and [assume]. Both commands are ways to indicate that a certain statement should hold any time this part of the program is reached. However they differ as follows: - If an [assert] statement fails, it causes the program to go into an error state and exit. - If an [assume] statement fails, the program fails to evaluate at all. In other words, the program gets stuck and has no final state. The new set of commands is: *) Module HoareAssertAssume. Inductive com : Type := | CSkip : com | CAsgn : string -> aexp -> com | CSeq : com -> com -> com | CIf : bexp -> com -> com -> com | CWhile : bexp -> com -> com | CAssert : bexp -> com | CAssume : bexp -> com. Notation "'assert' l" := (CAssert l) (in custom com at level 8, l custom com at level 0). Notation "'assume' l" := (CAssume l) (in custom com at level 8, l custom com at level 0). Notation "'skip'" := CSkip (in custom com at level 0). Notation "x := y" := (CAsgn x y) (in custom com at level 0, x constr at level 0, y at level 85, no associativity). Notation "x ; y" := (CSeq x y) (in custom com at level 90, right associativity). Notation "'if' x 'then' y 'else' z 'end'" := (CIf x y z) (in custom com at level 89, x at level 99, y at level 99, z at level 99). Notation "'while' x 'do' y 'end'" := (CWhile x y) (in custom com at level 89, x at level 99, y at level 99). (** To define the behavior of [assert] and [assume], we need to add notation for an error, which indicates that an assertion has failed. We modify the [ceval] relation, therefore, so that it relates a start state to either an end state or to [error]. The [result] type indicates the end value of a program, either a state or an error: *) Inductive result : Type := | RNormal : state -> result | RError : result. (** Now we are ready to give you the ceval relation for the new language. *) Inductive ceval : com -> state -> result -> Prop := (* Old rules, several modified *) | E_Skip : forall st, st =[ skip ]=> RNormal st | E_Asgn : forall st a1 n x, aeval st a1 = n -> st =[ x := a1 ]=> RNormal (x !-> n ; st) | E_SeqNormal : forall c1 c2 st st' r, st =[ c1 ]=> RNormal st' -> st' =[ c2 ]=> r -> st =[ c1 ; c2 ]=> r | E_SeqError : forall c1 c2 st, st =[ c1 ]=> RError -> st =[ c1 ; c2 ]=> RError | E_IfTrue : forall st r b c1 c2, beval st b = true -> st =[ c1 ]=> r -> st =[ if b then c1 else c2 end ]=> r | E_IfFalse : forall st r b c1 c2, beval st b = false -> st =[ c2 ]=> r -> st =[ if b then c1 else c2 end ]=> r | E_WhileFalse : forall b st c, beval st b = false -> st =[ while b do c end ]=> RNormal st | E_WhileTrueNormal : forall st st' r b c, beval st b = true -> st =[ c ]=> RNormal st' -> st' =[ while b do c end ]=> r -> st =[ while b do c end ]=> r | E_WhileTrueError : forall st b c, beval st b = true -> st =[ c ]=> RError -> st =[ while b do c end ]=> RError (* Rules for Assert and Assume *) | E_AssertTrue : forall st b, beval st b = true -> st =[ assert b ]=> RNormal st | E_AssertFalse : forall st b, beval st b = false -> st =[ assert b ]=> RError | E_Assume : forall st b, beval st b = true -> st =[ assume b ]=> RNormal st where "st '=[' c ']=>' r" := (ceval c st r). (** We redefine hoare triples: Now, [{{P}} c {{Q}}] means that, whenever [c] is started in a state satisfying [P], and terminates with result [r], then [r] is not an error and the state of [r] satisfies [Q]. *) Definition hoare_triple (P : Assertion) (c : com) (Q : Assertion) : Prop := forall st r, st =[ c ]=> r -> P st -> (exists st', r = RNormal st' /\ Q st'). Notation "{{ P }} c {{ Q }}" := (hoare_triple P c Q) (at level 90, c custom com at level 99) : hoare_spec_scope. (** To test your understanding of this modification, give an example precondition and postcondition that are satisfied by the [assume] statement but not by the [assert] statement. *) Theorem assert_assume_differ : exists (P:Assertion) b (Q:Assertion), ({{P}} assume b {{Q}}) /\ ~ ({{P}} assert b {{Q}}). (* FILL IN HERE *) Admitted. (** Then prove that any triple for an [assert] also works when [assert] is replaced by [assume]. *) Theorem assert_implies_assume : forall P b Q, ({{P}} assert b {{Q}}) -> ({{P}} assume b {{Q}}). Proof. (* FILL IN HERE *) Admitted. (** Next, here are proofs for the old hoare rules adapted to the new semantics. You don't need to do anything with these. *) Theorem hoare_asgn : forall Q X a, {{Q [X |-> a]}} X := a {{Q}}. Proof. unfold hoare_triple. intros Q X a st st' HE HQ. inversion HE. subst. exists (X !-> aeval st a ; st). split; try reflexivity. assumption. Qed. Theorem hoare_consequence_pre : forall (P P' Q : Assertion) c, {{P'}} c {{Q}} -> P ->> P' -> {{P}} c {{Q}}. Proof. intros P P' Q c Hhoare Himp. intros st st' Hc HP. apply (Hhoare st st'). assumption. apply Himp. assumption. Qed. Theorem hoare_consequence_post : forall (P Q Q' : Assertion) c, {{P}} c {{Q'}} -> Q' ->> Q -> {{P}} c {{Q}}. Proof. intros P Q Q' c Hhoare Himp. intros st r Hc HP. unfold hoare_triple in Hhoare. assert (exists st', r = RNormal st' /\ Q' st'). { apply (Hhoare st); assumption. } destruct H as [st' [Hr HQ'] ]. exists st'. split; try assumption. apply Himp. assumption. Qed. Theorem hoare_seq : forall P Q R c1 c2, {{Q}} c2 {{R}} -> {{P}} c1 {{Q}} -> {{P}} c1;c2 {{R}}. Proof. intros P Q R c1 c2 H1 H2 st r H12 Pre. inversion H12; subst. - eapply H1. + apply H6. + apply H2 in H3. apply H3 in Pre. destruct Pre as [st'0 [Heq HQ] ]. inversion Heq; subst. assumption. - (* Find contradictory assumption *) apply H2 in H5. apply H5 in Pre. destruct Pre as [st' [C _] ]. inversion C. Qed. (** Here are the other proof rules (sanity check) *) Theorem hoare_skip : forall P, {{P}} skip {{P}}. Proof. intros P st st' H HP. inversion H. subst. eexists. split. reflexivity. assumption. Qed. Theorem hoare_if : forall P Q (b:bexp) c1 c2, {{ P /\ b}} c1 {{Q}} -> {{ P /\ ~ b}} c2 {{Q}} -> {{P}} if b then c1 else c2 end {{Q}}. Proof. intros P Q b c1 c2 HTrue HFalse st st' HE HP. inversion HE; subst. - (* b is true *) apply (HTrue st st'). assumption. split. assumption. assumption. - (* b is false *) apply (HFalse st st'). assumption. split. assumption. apply bexp_eval_false. assumption. Qed. Theorem hoare_while : forall P (b:bexp) c, {{P /\ b}} c {{P}} -> {{P}} while b do c end {{ P /\ ~b}}. Proof. intros P b c Hhoare st st' He HP. remember <{while b do c end}> as wcom eqn:Heqwcom. induction He; try (inversion Heqwcom); subst; clear Heqwcom. - (* E_WhileFalse *) eexists. split. reflexivity. split. assumption. apply bexp_eval_false. assumption. - (* E_WhileTrueNormal *) clear IHHe1. apply IHHe2. reflexivity. clear IHHe2 He2 r. unfold hoare_triple in Hhoare. apply Hhoare in He1. + destruct He1 as [st1 [Heq Hst1] ]. inversion Heq; subst. assumption. + split; assumption. - (* E_WhileTrueError *) exfalso. clear IHHe. unfold hoare_triple in Hhoare. apply Hhoare in He. + destruct He as [st' [C _] ]. inversion C. + split; assumption. Qed. (** Finally, state Hoare rules for [assert] and [assume] and use them to prove a simple program correct. Name your rules [hoare_assert] and [hoare_assume]. *) (* FILL IN HERE *) (** Use your rules to prove the following triple. *) Example assert_assume_example: {{True}} assume (X = 1); X := X + 1; assert (X = 2) {{True}}. Proof. (* FILL IN HERE *) Admitted. End HoareAssertAssume. (** [] *)