MoreHoare: Hoare Logic with Lists; Formalizing Decorated Programs

(* Version of 8/24/2010 *)

Require Export SfLib.

This file introduces three generalizations of the material in the last few chapters:
  • A version of Imp (and Hoare logic) with list-manipulating primitives
  • A more "syntactic" treatment of the Hoare rules as a logical proof system in their own right
  • A formalization of the informal notion of "decorated programs" presented in Hoare.v

Hoare Logic with Lists

There are only so many simple, interesting functions that can be written using only numbers. In order to able to write a few more programs to reason about, we will introduce an extended version of Imp where variables can range over both numbers and lists of numbers. The basic operations will be extended to also include taking the head and tail of lists, and testing lists for nonemptyness.

Extensions


(*  In order to do this, we will only need to change the definitions
    of stateaexpaevalbexp, and beval. The definitions
    of comceval, and the Hoare logic rules (and their proofs)
    can be reused verbatim, although we need to copy-and-paste them in
    the context of the new definitions.  

    We start by repeating some material from Imp.v. *)


Inductive id : Type :=
  Id : nat -> id.

Definition beq_id id1 id2 :=
  match (id1, id2) with
    (Id n1, Id n2) => beq_nat n1 n2
  end.

Theorem beq_id_refl : forall i,
  true = beq_id i i.
Proof.
  intros. destruct i.
  apply beq_nat_refl. Qed.

Theorem beq_id_eq : forall i1 i2,
  true = beq_id i1 i2 -> i1 = i2.
Proof.
  intros i1 i2 H.
  destruct i1. destruct i2.
  apply beq_nat_eq in H. subst.
  reflexivity. Qed.

Theorem beq_id_false_not_eq : forall i1 i2,
  beq_id i1 i2 = false -> i1 <> i2.
Proof.
  intros i1 i2 H.
  destruct i1. destruct i2.
  apply beq_nat_false in H.
  intros C. apply H. inversion C. reflexivity. Qed.

Theorem not_eq_beq_id_false : forall i1 i2,
  i1 <> i2 -> beq_id i1 i2 = false.
Proof.
  intros i1 i2 H.
  destruct i1. destruct i2.
  assert (n <> n0).
    intros C. subst. apply H. reflexivity.
  apply not_eq_beq_false. assumption. Qed.

Definition X : id := Id 0.
Definition Y : id := Id 1.
Definition Z : id := Id 2.

Now we come to the key changes.
Rather than evaluating to a nat, an aexp in our new language will evaluate to a value val, which can be either a nat or a list of nats.
Similarly, states will now map identifiers to vals rather than nats, so that we can store lists in mutable variables.

Inductive val : Type :=
| VNat : nat -> val
| VList : list nat -> val.

Definition state := id -> val.

Definition empty_state : state := fun _ => VNat 0.
Definition update (st : state) (V:id) (n : val) : state :=
  fun V' => if beq_id V V' then n else st V'.

Imp does not have a type system, so nothing prevents the programmer from e.g. adding two lists or taking the head of a number. We have to decide what to do in such a (nonsensical) situation.
We adopt a simple solution: if an arithmetic function is given a list as an argument we treat the list as if it was the number 0. Similarly, if a list function is given a number as an argument we treat the number as if it was nil. (C.f. Javascript, where adding 3 to the empty list evaluates to 3...)
The two functions asnat and aslist interpret vals in a numeric or a list context; aeval calls these whenever it evaluates an arithmetic or a list operation.

Definition asnat (v : val) : nat :=
  match v with
  | VNat n => n
  | VList _ => 0
  end.

Definition aslist (v : val) : list nat :=
  match v with
  | VNat n => []
  | VList xs => xs
  end.

Now we fill in the definitions of abstract syntax and evaluation functions for arithmetic and boolean expressions.

Inductive aexp : Type :=
  | ANum : nat -> aexp
  | AId : id -> aexp
  | APlus : aexp -> aexp -> aexp
  | AMinus : aexp -> aexp -> aexp
  | AMult : aexp -> aexp -> aexp
  (* Four new cases: *)
  | AHead : aexp -> aexp
  | ATail : aexp -> aexp
  | ACons : aexp -> aexp -> aexp
  | ANil : aexp.

Tactic Notation "aexp_cases" tactic(first) tactic(c) :=
  first;
  [ c "ANum" | c "AId" | c "APlus" | c "AMinus" | c "AMult"
   | c "AHead" | c "ATail" | c "ACons" | c "ANil" ].

Definition tail (l : list nat) :=
  match l with
  | x::xs => xs
  | [] => []
  end.

Definition head (l : list nat) :=
  match l with
  | x::xs => x
  | [] => 0
  end.

Fixpoint aeval (st : state) (e : aexp) : val :=
  match e with
  | ANum n => VNat n
  | AId i => st i
  | APlus a1 a2 => VNat (asnat (aeval st a1) + asnat (aeval st a2))
  | AMinus a1 a2 => VNat (asnat (aeval st a1) - asnat (aeval st a2))
  | AMult a1 a2 => VNat (asnat (aeval st a1) * asnat (aeval st a2))
  (* Four new cases: *)
  | ATail a => VList (tail (aslist (aeval st a)))
  | AHead a => VNat (head (aslist (aeval st a)))
  | ACons a1 a2 => VList (asnat (aeval st a1) :: aslist (aeval st a2))
  | ANil => VList []
  end.

We extend bexps with an operation to test if a list is nonempty, and adapt beval acordingly.

Inductive bexp : Type :=
  | BTrue : bexp
  | BFalse : bexp
  | BEq : aexp -> aexp -> bexp
  | BLe : aexp -> aexp -> bexp
  | BNot : bexp -> bexp
  | BAnd : bexp -> bexp -> bexp
  (* New case: *)
  | BIsCons : aexp -> bexp.

Tactic Notation "bexp_cases" tactic(first) tactic(c) :=
  first;
  [ c "BTrue" | c "BFalse" |
    c "BEq" | c "BLe" |
    c "BNot" | c "BAnd" |
    c "BIsCons" ].

Fixpoint beval (st : state) (e : bexp) : bool :=
  match e with
  | BTrue => true
  | BFalse => false
  | BEq a1 a2 => beq_nat (asnat (aeval st a1)) (asnat (aeval st a2))
  | BLe a1 a2 => ble_nat (asnat (aeval st a1)) (asnat (aeval st a2))
  | BNot b1 => negb (beval st b1)
  | BAnd b1 b2 => andb (beval st b1) (beval st b2)
  (* New case: *)
  | BIsCons a => match aslist (aeval st a) with
                     | _::_ => true
                     | [] => false
                   end
  end.

Repeated Definitions

We can keep exactly the same old definitions of com, ceval, Assertion, and hoare_triple, and the same proofs of all the rules of Hoare logic.

Inductive com : Type :=
  | CSkip : com
  | CAss : id -> aexp -> com
  | CSeq : com -> com -> com
  | CIf : bexp -> com -> com -> com
  | CWhile : bexp -> com -> com.

Tactic Notation "com_cases" tactic(first) tactic(c) :=
  first;
  [ c "SKIP" | c "::=" | c ";" | c "IFB" | c "WHILE" ].

Notation "'SKIP'" :=
  CSkip (at level 10).
Notation "l '::=' a" :=
  (CAss l a) (at level 60).
Notation "c1 ; c2" :=
  (CSeq c1 c2) (at level 80, right associativity).
Notation "'WHILE' b 'DO' c 'END'" :=
  (CWhile b c) (at level 80, right associativity).
Notation "'IFB' e1 'THEN' e2 'ELSE' e3 'FI'" :=
  (CIf e1 e2 e3) (at level 80, right associativity).

Reserved Notation "c1 '/' st '==>' st'" (at level 40).

Inductive ceval : state -> com -> state -> Prop :=
  | E_Skip : forall st,
      SKIP / st ==> st
  | E_Asgn : forall st a1 n l,
      aeval st a1 = n ->
      (l ::= a1) / st ==> (update st l n)
  | E_Seq : forall c1 c2 st st' st'',
      c1 / st ==> st' ->
      c2 / st' ==> st'' ->
      (c1 ; c2) / st ==> st''
  | E_IfTrue : forall st st' b1 c1 c2,
      beval st b1 = true ->
      c1 / st ==> st' ->
      (IFB b1 THEN c1 ELSE c2 FI) / st ==> st'
  | E_IfFalse : forall st st' b1 c1 c2,
      beval st b1 = false ->
      c2 / st ==> st' ->
      (IFB b1 THEN c1 ELSE c2 FI) / st ==> st'
  | E_WhileEnd : forall b1 st c1,
      beval st b1 = false ->
      (WHILE b1 DO c1 END) / st ==> st
  | E_WhileLoop : forall st st' st'' b1 c1,
      beval st b1 = true ->
      c1 / st ==> st' ->
      (WHILE b1 DO c1 END) / st' ==> st'' ->
      (WHILE b1 DO c1 END) / st ==> st''

  where "c1 '/' st '==>' st'" := (ceval st c1 st').

Tactic Notation "ceval_cases" tactic(first) tactic(c) := first; [
    c "E_Skip" | c "E_Asgn" | c "E_Seq" | c "E_IfTrue" | c "E_IfFalse"
  | c "E_WhileEnd" | c "E_WhileLoop" ].

Theorem update_eq : forall n V st,
  (update st V n) V = n.
Proof.
  intros n V st.
  unfold update.
  rewrite <- beq_id_refl.
  reflexivity.
Qed.

Theorem update_neq : forall V2 V1 n st,
  beq_id V2 V1 = false ->
  (update st V2 n) V1 = (st V1).
Proof.
  intros V2 V1 n st Hneq.
  unfold update.
  rewrite -> Hneq.
  reflexivity. Qed.

Theorem update_shadow : forall x1 x2 k1 k2 (f : state),
   (update (update f k2 x1) k2 x2) k1 = (update f k2 x2) k1.
Proof.
  intros x1 x2 k1 k2 f.
  unfold update.
  destruct (beq_id k2 k1); reflexivity. Qed.

Theorem update_same : forall x1 k1 k2 (f : state),
  f k1 = x1 ->
  (update f k1 x1) k2 = f k2.
Proof.
  intros x1 k1 k2 f Heq.
  unfold update. subst.
  remember (beq_id k1 k2) as b.
  destruct b.
  Case "true".
    apply beq_id_eq in Heqb. subst. reflexivity.
  Case "false".
    reflexivity. Qed.

Theorem update_permute : forall x1 x2 k1 k2 k3 f,
  beq_id k2 k1 = false ->
  (update (update f k2 x1) k1 x2) k3 = (update (update f k1 x2) k2 x1) k3.
Proof.
  intros x1 x2 k1 k2 k3 f H.
  unfold update.
  remember (beq_id k1 k3) as b13.
  remember (beq_id k2 k3) as b23.
  apply beq_id_false_not_eq in H.
  destruct b13; try reflexivity.
  Case "true".
    destruct b23; try reflexivity.
    SCase "true".
      apply beq_id_eq in Heqb13.
      apply beq_id_eq in Heqb23.
      subst. apply ex_falso_quodlibet. apply H. reflexivity. Qed.

Definition Assertion := state -> Prop.

Definition hoare_triple (P:Assertion) (c:com) (Q:Assertion) : Prop :=
  forall st st',
       c / st ==> st' ->
       P st ->
       Q st'.

Notation "{{ P }} c {{ Q }}" := (hoare_triple P c Q) (at level 90).

Definition assn_sub V a Q : Assertion :=
  fun (st : state) =>
    Q (update st V (aeval st a)).

Theorem hoare_asgn : forall Q V a,
  {{assn_sub V a Q}} (V ::= a) {{Q}}.
Proof.
  unfold hoare_triple.
  intros Q V a st st' HE HQ.
  inversion HE. subst.
  unfold assn_sub in HQ. assumption. Qed.

Theorem hoare_asgn_eq : forall Q Q' V a,
     Q' = assn_sub V a Q ->
     {{Q'}} (V ::= a) {{Q}}.
Proof.
  intros Q Q' V a H. rewrite H. apply hoare_asgn. Qed.

Theorem hoare_consequence : forall (P P' Q Q' : Assertion) c,
  {{P'}} c {{Q'}} ->
  (forall st, P st -> P' st) ->
  (forall st, Q' st -> Q st) ->
  {{P}} c {{Q}}.
Proof.
  unfold hoare_triple.
  intros P P' Q Q' c Hht HPP' HQ'Q.
  intros st st' Hc HP.
  apply HQ'Q. apply (Hht st st'). assumption.
  apply HPP'. assumption. Qed.

Theorem hoare_consequence_pre : forall (P P' Q : Assertion) c,
  {{P'}} c {{Q}} ->
  (forall st, P st -> P' st) ->
  {{P}} c {{Q}}.
Proof.
  intros P P' Q c Hhoare Himp.
  apply hoare_consequence with (P' := P') (Q' := Q);
    try assumption.
  intros st H. apply H. Qed.

Theorem hoare_consequence_post : forall (P Q Q' : Assertion) c,
  {{P}} c {{Q'}} ->
  (forall st, Q' st -> Q st) ->
  {{P}} c {{Q}}.
Proof.
  intros P Q Q' c Hhoare Himp.
  apply hoare_consequence with (P' := P) (Q' := Q');
    try assumption.
  intros st H. apply H. Qed.

Theorem hoare_skip : forall P,
     {{P}} SKIP {{P}}.
Proof.
  unfold hoare_triple. intros P st st' H HP. inversion H. subst.
  assumption. Qed.

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.
  apply (H1 st'0 st'); try assumption.
  apply (H2 st st'0); try assumption. Qed.

Definition bassn b : Assertion :=
  fun st => beval st b = true.

Lemma bexp_eval_true : forall b st,
  beval st b = true -> (bassn b) st.
Proof.
  intros b st Hbe.
  unfold bassn. assumption. Qed.

Lemma bexp_eval_false : forall b st,
  beval st b = false -> ~ ((bassn b) st).
Proof.
  intros b st Hbe contra.
  unfold bassn in contra.
  rewrite -> contra in Hbe. inversion Hbe. Qed.

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}} (IFB b THEN c1 ELSE c2 FI) {{Q}}.
Proof.
  unfold hoare_triple.
  intros P Q b c1 c2 HTrue HFalse st st' HE HP.
  inversion HE; subst.
  Case "b is true".
    apply (HTrue st st').
      assumption.
      split. assumption.
             apply bexp_eval_true. assumption.
  Case "b is false".
    apply (HFalse st st').
      assumption.
      split. assumption.
             apply bexp_eval_false. assumption. Qed.

Lemma hoare_while : forall P b c,
  {{fun st => P st /\ bassn b st}} c {{P}} ->
  {{P}} (WHILE b DO c END) {{fun st => P st /\ ~ (bassn b st)}}.
Proof.
  unfold hoare_triple.
  intros P b c Hhoare st st' He HP.
  remember (WHILE b DO c END) as wcom.
  (ceval_cases (induction He) Case); try (inversion Heqwcom); subst.

  Case "E_WhileEnd".
    split. assumption. apply bexp_eval_false in H. apply H.

  Case "E_WhileLoop".
    destruct IHHe2 as [HP' Hbfalse']; try assumption; try reflexivity.
    apply (Hhoare st st'); try assumption.
      split. assumption. apply bexp_eval_true. assumption.
    split; assumption. Qed.

Examples

When writing down assertions formally we'll need to use the functions asnat and aslist to "cast" values to the right type. When we write the assertions informally we elide the casts. For example, the informal proof
                    {{ X + 1 <= 5 }}
      X ::= X + 1
                    {{ X <= 5 }}
corresponds to the following formal proof:

Example assn_sub_ex :
  {{ fun st => asnat (st X) + 1 <= 5 }}
  (X ::= APlus (AId X) (ANum 1))
  {{ fun st => asnat (st X) <= 5 }}.
Proof.
  apply hoare_asgn_eq. reflexivity. Qed.

Exercise: 3 stars (list_sum)

Here is a direct definition of the sum of the elements of a list, and an Imp program that computes the sum.

Definition sum l := fold_right plus 0 l.

Definition sum_program :=
  Y ::= ANum 0;
  WHILE (BIsCons (AId X)) DO
    Y ::= APlus (AId Y) (AHead (AId X)) ;
    X ::= ATail (AId X)
  END.

Provide an informal proof of the following specification of sum_program in the form of a decorated version of the program.

Definition sum_program_spec := forall l,
  {{ fun st => aslist (st X) = l }}
  sum_program
  {{ fun st => asnat (st Y) = sum l }}.

(* FILL IN HERE *)
Next, let's look at a formal Hoare Logic proof for a program that deals with lists. We will verify the following program, which checks if the number Y occurs in the list X, and if so sets Z to 1.

Definition list_member :=
  WHILE BIsCons (AId X) DO
    IFB (BEq (AId Y) (AHead (AId X))) THEN
      Z ::= (ANum 1)
    ELSE
      SKIP
    FI;
    X ::= ATail (AId X)
  END.

In order to state its specification, we define a relation appears_in, that declaratively specifies what it means for an element to appear in a list.

Inductive appears_in (n : nat) : list nat -> Prop :=
| ai_here : forall l, appears_in n (n::l)
| ai_later : forall m l, appears_in n l -> appears_in n (m::l).

Definition list_member_spec := forall l n,
  {{ fun st => st X = VList l /\ st Y = VNat n /\ st Z = VNat 0 }}
  list_member
  {{ fun st => st Z = VNat 1 <-> appears_in n l }}.

The proof we will use, written informally, looks as follows:
                   {{ X = l /\ Y = n /\ Z = 0 }}
                   =>
                   {{ Y = n
                     /\ exists p, p ++ X = l
                               /\ (Z = 1 <-> appears_in n p) }}
  WHILE (BIsCons X)
  DO
                   {{ Y = n
                      /\ (exists p, p ++ X = l
                                 /\ (Z = 1 <-> appears_in n p))
                      /\ (BIsCons X) }}
    IFB (Y == head X) THEN
                   {{ Y = n
                       /\ (exists p, p ++ X = l
                                  /\ (Z = 1 <-> appears_in n p))
                       /\ (BIsCons X)
                       /\ Y == AHead X }}
                       =>
                    {{ Y = n
                       /\ (exists p, p ++ tail X = l
                                  /\ (1 = 1 <-> appears_in n p)) }}
      Z ::= 1
                    {{ Y = n
                       /\ (exists p, p ++ tail X = l
                                  /\ (Z = 1 <-> appears_in n p)) }}
    ELSE
                   {{ Y = n
                      /\ (exists p, p ++ X = l
                                 /\ (Z = 1 <-> appears_in n p))
                      /\ (BIsCons X)
                      /\ ~ (Y == head X) }}
                   =>
                   {{ Y = n
                      /\ (exists p, p ++ tail X = l
                                 /\ (Z = 1 <-> appears_in n p)) }}
      SKIP
                   {{ Y = n
                      /\ (exists p, p ++ tail X = l
                                /\ (Z = 1 <-> appears_in n p)) }}
    FI;
    X ::= ATail X
                  {{ Y = n
                     /\ (exists p, p ++ X = l
                                /\ (Z = 1 <-> appears_in n p)) }}
  END
                  {{ Y = n
                     /\ (exists p, p ++ X = l
                                /\ (Z = 1 <-> appears_in n p))
                     /\ ~ (BIsCons X) }}
                  =>
                  {{ Z = 1 <-> appears_in n l }}
The only interesting part of the proof is the choice of loop invariant:
      exists p, p ++ X = l /\ (Z = 1 <-> appears_in n p)
This states that at each iteration of the loop, the original list l is equal to the append of the current value of X and some other list p which is not the value of any variable in the program, but keeps track of enough information from the original state to make the proof go through. (Such a p is sometimes called a "ghost variable").
In order to show that such a list p exists, in each iteration we add the head of X to the end of p. This needs the function snoc, from Poly.v.

Fixpoint snoc {X:Type} (l:list X) (v:X) : (list X) :=
  match l with
  | nil => [ v ]
  | cons h t => h :: (snoc t v)
  end.

Lemma snoc_equation : forall (A : Type) (h:A) (x y : list A),
  snoc x h ++ y = x ++ h :: y.
Proof.
  intros A h x y.
  induction x.
    reflexivity.
    simpl. rewrite IHx. reflexivity.
Qed.

The main proof uses various lemmas.

Lemma appears_in_snoc1 : forall a l,
  appears_in a (snoc l a).
Proof.
  induction l.
    apply ai_here.
    simpl. apply ai_later. apply IHl.
Qed.

Lemma appears_in_snoc2 : forall a b l,
  appears_in a l ->
  appears_in a (snoc l b).
Proof.
  induction l; inversion 1; subst; simpl.
    apply ai_here.
    apply ai_later. apply IHl. assumption.
Qed.

Lemma appears_in_snoc3 : forall a b l,
   appears_in a (snoc l b) ->
   (appears_in a l \/ a = b).
Proof.
   induction l.
     inversion 1.
       right. reflexivity.
       inversion H1.
     inversion 1; subst.
       left. apply ai_here.
       destruct (IHl H1).
         left. apply ai_later. assumption.
         right. assumption.
Qed.

Lemma append_singleton_equation : forall (x : nat) l l',
  (l ++ [x]) ++ l' = l ++ x :: l'.
Proof.
  intros x l l'.
  induction l.
    reflexivity.
    simpl. rewrite IHl. reflexivity.
Qed.

Lemma append_nil : forall (A : Type) (l : list A),
  l ++ [] = l.
Proof.
  induction l.
    reflexivity.
    simpl. rewrite IHl. reflexivity.
Qed.

Lemma beq_true__eq : forall n n',
  beq_nat n n' = true ->
  n = n'.
Proof.
 induction n.
   destruct n'.
     reflexivity.
     simpl. intros H. inversion H.
   destruct n'.
     simpl. intros H. inversion H.
     simpl. intros H.
       rewrite (IHn n' H). reflexivity.
Qed.

Lemma beq_nat_refl : forall n,
  beq_nat n n = true.
Proof.
  induction n.
    reflexivity.
    simpl. assumption.
Qed.

Theorem list_member_correct : forall l n,
  {{ fun st => st X = VList l /\ st Y = VNat n /\ st Z = VNat 0 }}
  list_member
  {{ fun st => st Z = VNat 1 <-> appears_in n l }}.
Proof.
  intros l n.
  eapply hoare_consequence.
  apply hoare_while with (P := fun st =>
     st Y = VNat n
     /\ exists p, p ++ aslist (st X) = l
                  /\ (st Z = VNat 1 <-> appears_in n p)).
    (* The loop body preserves the invariant: *)
    eapply hoare_seq.
    apply hoare_asgn.
    apply hoare_if.
    Case "If taken".
      eapply hoare_consequence_pre.
      apply hoare_asgn.
      intros st [[[H1 [p [H2 H3]]] H9] H10].
      unfold assn_sub. split.
      (* (st Y) is still n *)
        rewrite update_neq; try reflexivity.
        rewrite update_neq; try reflexivity.
        assumption.
      (* and the interesting part of the invariant is preserved: *)
        (* X has to be a cons *)
        remember (aslist (st X)) as x.
        destruct x as [|h x'].
          unfold bassn in H9. unfold beval in H9. unfold aeval in H9. rewrite <- Heqx in H9.
          inversion H9.

          exists (snoc p h).
          rewrite update_eq.
          unfold aeval. rewrite update_neq; try reflexivity.
          rewrite <- Heqx.
          split.
            rewrite snoc_equation. assumption.

            rewrite update_neq; [idtac | reflexivity].
            rewrite update_eq.
            split.
              simpl.
              unfold bassn in H10. unfold beval in H10.
              unfold aeval in H10. rewrite H1 in H10. rewrite <- Heqx in H10.
              simpl in H10.
              rewrite (beq_true__eq _ _ H10).
              intros. apply appears_in_snoc1.

              intros. reflexivity.
    Case "If not taken".
      eapply hoare_consequence_pre. apply hoare_skip.
      unfold assn_sub.
      intros st [[[H1 [p [H2 H3]]] H9] H10].
      split.
      (* (st Y) is still n *)
        rewrite update_neq; try reflexivity.
        assumption.
      (* and the interesting part of the invariant is preserved: *)
        (* X has to be a cons *)
        remember (aslist (st X)) as x.
        destruct x as [|h x'].
          unfold bassn in H9. unfold beval in H9. unfold aeval in H9. rewrite <- Heqx in H9.
          inversion H9.

          exists (snoc p h).
          split.
            rewrite update_eq.
            unfold aeval. rewrite <- Heqx.
            rewrite snoc_equation. assumption.

            rewrite update_neq; [idtac | reflexivity].
            split.
              intros. apply appears_in_snoc2. apply H3. assumption.

              intros. destruct (appears_in_snoc3 _ _ _ H).
              SCase "later".
                destruct H3 as [H3 H3'].
                apply H3'. assumption.
              SCase "here (absurd)".
                subst.
                unfold bassn in H10. unfold beval in H10. unfold aeval in H10.
                rewrite <- Heqx in H10. rewrite H1 in H10.
                simpl in H10. rewrite beq_nat_refl in H10.
                apply ex_falso_quodlibet. apply H10. reflexivity.
  (* The invariant holds at the start of the loop: *)
  intros st [H1 [H2 H3]].
  rewrite H1. rewrite H2. rewrite H3.
  split.
    reflexivity.
    exists []. split.
      reflexivity.
      split; intros H; inversion H.
  (* At the end of the loop the invariant implies the right thing. *)
  simpl. intros st [[H1 [p [H2 H3]]] H5].
  (* x must be  *)
  unfold bassn in H5. unfold beval in H5. unfold aeval in H5.
  destruct (aslist (st X)) as [|h x'].
    rewrite append_nil in H2.
    rewrite <- H2.
    assumption.

    apply ex_falso_quodlibet. apply H5. reflexivity.
Qed.

Exercise: 4 stars (list_reverse)

Recall the function rev from Poly.v, for reversing lists.

Fixpoint rev {X:Type} (l:list X) : list X :=
  match l with
  | nil => []
  | cons h t => snoc (rev t) h
  end.

Write an Imp program list_reverse_program that reverses lists. Formally prove that it satisfies the following specification:
    forall l : list nat,
      {{ X = l /\ Y = nil }}
      list_reverse_program
      {{ Y = rev l }}.
You may find the lemmas append_nil and rev_equation useful.

Lemma rev_equation : forall (A : Type) (h : A) (x y : list A),
  rev (h :: x) ++ y = rev x ++ h :: y.
Proof.
  intros. simpl. apply snoc_equation.
Qed.

(* FILL IN HERE *)

Hoare Logic as a Logic

The presentation of Hoare logic in Hoare.v could be described as "model-theoretic": the proof rules for each of the constructors were presented as theorems about the evaluation behavior of programs, and proofs of program correctness (validity of Hoare triples) were constructed by combining these theorems directly in Coq.
Another way of presenting Hoare logic is to define a completely separate proof system -- a set of axioms and inference rules that talk about commands, Hoare triples, etc. -- and then say that a proof of a Hoare triple is a valid derivation in that logic. We can do this by giving an inductive definition of valid derivations in this new logic.

Inductive hoare_proof : Assertion -> com -> Assertion -> Type :=
  | H_Skip : forall P,
      hoare_proof P (SKIP) P
  | H_Asgn : forall Q V a,
      hoare_proof (assn_sub V a Q) (V ::= a) Q
  | H_Seq : forall P c Q d R,
      hoare_proof P c Q -> hoare_proof Q d R -> hoare_proof P (c;d) R
  | H_If : forall P Q b c1 c2,
    hoare_proof (fun st => P st /\ bassn b st) c1 Q ->
    hoare_proof (fun st => P st /\ ~(bassn b st)) c2 Q ->
    hoare_proof P (IFB b THEN c1 ELSE c2 FI) Q
  | H_While : forall P b c,
    hoare_proof (fun st => P st /\ bassn b st) c P ->
    hoare_proof P (WHILE b DO c END) (fun st => P st /\ ~ (bassn b st))
  | H_Consequence : forall (P Q P' Q' : Assertion) c,
    hoare_proof P' c Q' ->
    (forall st, P st -> P' st) ->
    (forall st, Q' st -> Q st) ->
    hoare_proof P c Q
  | H_Consequence_pre : forall (P Q P' : Assertion) c,
    hoare_proof P' c Q ->
    (forall st, P st -> P' st) ->
    hoare_proof P c Q
  | H_Consequence_post : forall (P Q Q' : Assertion) c,
    hoare_proof P c Q' ->
    (forall st, Q' st -> Q st) ->
    hoare_proof P c Q.

Tactic Notation "hoare_proof_cases" tactic(first) tactic(c) :=
  first;
  [ c "H_Skip" | c "H_Asgn" | c "H_Seq" | c "H_If" | c "H_While" |
    c "H_Consequence" | c "H_Consequence_pre" | c "H_Consequence_post" ].

For example, let's construct a proof object representing a derivation for the hoare triple
      {{assn_sub X (X+1) (assn_sub X (X+2) (X=3))}} X::=X+1; X::=X+2 {{X=3}}.
We can use Coq's tactics to help us construct the proof object.

Example sample_proof
             : hoare_proof
                 (assn_sub X (APlus (AId X) (ANum 1))
                   (assn_sub X (APlus (AId X) (ANum 2))
                     (fun st => st X = VNat 3) ))
                 (X ::= APlus (AId X) (ANum 1); (X ::= APlus (AId X) (ANum 2)))
                 (fun st => st X = VNat 3).
Proof.
  apply H_Seq with (assn_sub X (APlus (AId X) (ANum 2))
                     (fun st => st X = VNat 3)).
  apply H_Asgn. apply H_Asgn.
Qed.

(* 
Print sample_proof.
====> 
  H_Seq
    (assn_sub X (APlus (AId X) (ANum 1))
       (assn_sub X (APlus (AId X) (ANum 2)) (fun st : state => st X = VNat 3)))
    (X ::= APlus (AId X) (ANum 1))
    (assn_sub X (APlus (AId X) (ANum 2)) (fun st : state => st X = VNat 3))
    (X ::= APlus (AId X) (ANum 2)) (fun st : state => st X = VNat 3)
    (H_Asgn
       (assn_sub X (APlus (AId X) (ANum 2)) (fun st : state => st X = VNat 3))
       X (APlus (AId X) (ANum 1)))
    (H_Asgn (fun st : state => st X = VNat 3) X (APlus (AId X) (ANum 2)))
*)


Exercise: 2 stars

Prove that such proof objects represent true claims.

Theorem hoare_proof_sound : forall P c Q,
  hoare_proof P c Q -> {{P}} c {{Q}}.
Proof.
  (* FILL IN HERE *) Admitted.
We can also use Coq's reasoning facilities to prove metatheorems about Hoare Logic. For example, here are the analogs of two theorems we saw in Hoare.v -- this time expressed in terms of the syntax of Hoare Logic derivations (provability) rather than directly in terms of the semantics of Hoare triples.
The first one says that, for every P and c, the assertion {{P}} c {{True}} is provable in Hoare Logic. Note that the proof is more complex than the semantic proof in Hoare.v: we actually need to perform an induction over the structure of the command c.

Theorem H_Post_True_deriv:
  forall c P, hoare_proof P c (fun _ => True).
Proof.
  intro c.
  (com_cases (induction c) Case); intro P.
  Case "SKIP".
    eapply H_Consequence_pre.
    apply H_Skip.
    (* Proof of True *)
    intros. apply I.
  Case "::=".
    eapply H_Consequence_pre.
    apply H_Asgn.
    intros. apply I.
  Case ";".
    eapply H_Consequence_pre.
    eapply H_Seq.
    apply (IHc1 (fun _ => True)).
    apply IHc2.
    intros. apply I.
  Case "IFB".
    apply H_Consequence_pre with (fun _ => True).
    apply H_If.
    apply IHc1.
    apply IHc2.
    intros. apply I.
  Case "WHILE".
    eapply H_Consequence.
    eapply H_While.
    eapply IHc.
    intros; apply I.
    intros; apply I.
Qed.

Similarly, we can show that {{False}} c {{Q}} is provable for any c and Q.

Lemma False_and_P_imp: forall P Q,
  False /\ P -> Q.
Proof.
  intros P Q [CONTRA HP].
  destruct CONTRA.
Qed.

Tactic Notation "pre_false_helper" constr(CONSTR) :=
  eapply H_Consequence_pre;
    [eapply CONSTR | intros ? CONTRA; destruct CONTRA].

Theorem H_Pre_False_deriv:
  forall c Q, hoare_proof (fun _ => False) c Q.
Proof.
  intros c.
  (com_cases (induction c) Case); intro Q.
  Case "SKIP". pre_false_helper H_Skip.
  Case "::=". pre_false_helper H_Asgn.
  Case ";". pre_false_helper H_Seq. apply IHc1. apply IHc2.
  Case "IFB".
    apply H_If; eapply H_Consequence_pre.
    apply IHc1. intro. eapply False_and_P_imp.
    apply IHc2. intro. eapply False_and_P_imp.
  Case "WHILE".
    eapply H_Consequence_post.
    eapply H_While.
    eapply H_Consequence_pre.
      apply IHc.
      intro. eapply False_and_P_imp.
    intro. simpl. eapply False_and_P_imp.
Qed.

This style of presentation gives a clearer picture of what it means to "give a proof in Hoare logic." However, it is not entirely satisfactory from the point of view of writing down such proofs in practice: it is quite verbose. The next section shows how we can do even better.

Formalizing Decorated Programs

In Hoare.v we developed a set of informal rules for displaying Hoare triples in which commands were annotated with enough embedded assertions that checking the validity of the triple was reduced to simple algebraic calculations showing that some assertions were stronger than others.
In this section, we show that this informal presentation style can actually be made completely formal.

Syntax

The first thing we need to do is to formalize a variant of the syntax of commands that includes embedded assertions. We call the new commands decorated commands, or dcoms.

Inductive dcom : Type :=
  | DCSkip : Assertion -> dcom
  | DCSeq : dcom -> dcom -> dcom
  | DCAsgn : id -> aexp -> Assertion -> dcom
  | DCIf : bexp -> Assertion -> dcom -> Assertion -> dcom -> dcom
  | DCWhile : bexp -> Assertion -> dcom -> dcom
  | DCPre : Assertion -> dcom -> dcom
  | DCPost : dcom -> Assertion -> dcom.

Tactic Notation "dcom_cases" tactic(first) tactic(c) :=
  first;
  [ c "Skip" | c "Seq" | c "Asgn" | c "If" | c "While"
    | c "Pre" | c "Post" ].

Notation "'SKIP' {{ P }}"
      := (DCSkip P)
      (at level 10) : dcom_scope.
Notation "l '::=' a {{ P }}"
      := (DCAsgn l a P)
      (at level 60, a at next level) : dcom_scope.
Notation "'WHILE' b 'DO' {{ P }} d 'END'"
      := (DCWhile b P d)
      (at level 80, right associativity) : dcom_scope.
Notation "'IFB' b 'THEN' {{ P }} d 'ELSE' {{ P' }} d' 'FI'"
      := (DCIf b P d P' d')
      (at level 80, right associativity) : dcom_scope.
Notation "'=>' {{ P }} d"
      := (DCPre P d)
      (at level 90, right associativity) : dcom_scope.
Notation "d '=>' {{ P }}"
      := (DCPost d P)
      (at level 91, right associativity) : dcom_scope.
Notation " d ; d' "
      := (DCSeq d d')
      (at level 80, right associativity) : dcom_scope.

Delimit Scope dcom_scope with dcom.

To avoid clashing with the existing Notation definitions for ordinary commands, we introduce these notations in a special scope called dcom_scope, and we wrap examples with the declaration % dcom to signal that we want the notations to be interpreted in this scope.

Example dec_while : dcom := (
  WHILE (BNot (BEq (AId X) (ANum 0)))
  DO
                             {{ fun _ => True }}
      X ::= (AMinus (AId X) (ANum 1))
                             {{ fun _ => True }}
  END
                             =>
                             {{ fun st => asnat (st X) = 0 }}

) % dcom.

It is easy to go from a dcom to a com by erasing all annotations.

Fixpoint extract (d:dcom) : com :=
  match d with
  | DCSkip _ => SKIP
  | DCSeq c1 c2 => (extract c1 ; extract c2)
  | DCAsgn V a _ => V ::= a
  | DCIf b _ t _ e => IFB b THEN extract t ELSE extract e FI
  | DCWhile b _ c => WHILE b DO extract c END
  | DCPre _ c => extract c
  | DCPost c _ => extract c
  end.

The choice of exactly where to put assertions in the definition of dcom is a bit subtle. The simplest thing to do would be to annotate every dcom with a precondition and postcondition. But this would result in very verbose programs with a lot of repeated annotations: for example, a program like SKIP;SKIP would have to be annotated as
        {{P}} ({{P}} SKIP {{P}}) ; ({{P}} SKIP {{P}}) {{P}},
with pre- and post-conditions on each SKIP, plus identical pre- and post-conditions on the semicolon!
Instead, the rule we've followed is this:
  • The post-condition expected by each dcom d is embedded in d
  • The pre-condition is supplied by the context.
In other words, the invariant of the representation is that a dcom d together with a precondition P determines a Hoare triple {{P}} (extract d) {{post d}}, where post is defined as follows:

Fixpoint post (d:dcom) : Assertion :=
  match d with
  | DCSkip P => P
  | DCSeq c1 c2 => post c2
  | DCAsgn V a Q => Q
  | DCIf _ _ t _ e => post t
  | DCWhile b _ c => (fun st => post c st /\ ~ bassn b st)
  | DCPre _ c => post c
  | DCPost c Q => Q
  end.

Definition dec_correct (P : Assertion) (d:dcom) :=
  {{P}} (extract d) {{post d}}.

To check whether this Hoare triple is valid, we need a way to extract the "proof obligations" from a decorated program. These obligations are often called verification conditions, because they are the facts that must be verified (by some process looking at the decorated program) to see that the decorations are logically consistent and thus add up to a proof of correctness.

Extracting Verification Conditions

First, a bit of notation:

Definition assert_implies (P Q : Assertion) : Prop :=
  forall st, P st -> Q st.

Notation "P ~~> Q" := (assert_implies P Q) (at level 80).

Next, the key definition. The function verification_conditions takes a dcom d together with a precondition P and returns a proposition that, if it can be proved, implies that the triple {{P}} (extract d) {{post d}} is valid. It does this by walking over d and generating a big conjunction including all the "local checks" that we listed when we described the informal rules for decorated programs. (Strictly speaking, we need to massage the informal rules a little bit to add some uses of the rule of consequence, but the correspondence should be clear.)

Fixpoint verification_conditions (P : Assertion) (d:dcom) : Prop :=
  match d with
  | DCSkip Q => (P ~~> Q)
  | DCSeq c1 c2 => verification_conditions P c1
                      /\ verification_conditions (post c1) c2
  | DCAsgn V a Q => (P ~~> assn_sub V a Q)
  | DCIf b P1 t P2 e => ((fun st => P st /\ bassn b st) ~~> P1)
                      /\ ((fun st => P st /\ ~ (bassn b st)) ~~> P2)
                      /\ (post t = post e)
                      /\ verification_conditions P1 t
                      /\ verification_conditions P2 e
  | DCWhile b P' c => ((fun st => post c st /\ bassn b st) ~~> P')
                      /\ (P ~~> post c) (* post c is the loop invariant *)
                      /\ verification_conditions P' c
  | DCPre P' c => (P ~~> P') /\ verification_conditions P' c
  | DCPost c Q => verification_conditions P c /\ (post c ~~> Q)
  end.

And now, the key theorem, which captures the claim that the verification_conditions function does its job correctly. Not surprisingly, we need all of the Hoare Logic rules in the proof.

Theorem verification_correct : forall d P,
  verification_conditions P d -> dec_correct P d.
Proof.
  (dcom_cases (induction d) Case); intros P H; simpl in *.
  Case "Skip".
    eapply hoare_consequence_pre.
      apply hoare_skip.
      assumption.
  Case "Seq".
    inversion H as [H1 H2]. clear H.
    eapply hoare_seq.
      apply IHd2. apply H2.
      apply IHd1. apply H1.
  Case "Asgn".
    eapply hoare_consequence_pre.
      apply hoare_asgn.
      assumption.
  Case "If".
    inversion H as [HPre1 [HPre2 [HQeq [HThen HElse]]]]; clear H.
    apply hoare_if.
      eapply hoare_consequence_pre. apply IHd1. apply HThen. assumption.
      simpl. rewrite HQeq.
      eapply hoare_consequence_pre. apply IHd2. apply HElse. assumption.
  Case "While".
    inversion H as [HPre [HPreBody Hd]]; clear H.
    eapply hoare_consequence.
      apply hoare_while with (P := post d).
        eapply hoare_consequence_pre. apply IHd. apply Hd. assumption.
      assumption.
      simpl. intros. assumption.
  Case "Pre".
    inversion H as [HP Hd]; clear H.
    eapply hoare_consequence_pre. apply IHd. apply Hd. assumption.
  Case "Post".
    inversion H as [Hd HQ]; clear H.
    eapply hoare_consequence_post. apply IHd. apply Hd. assumption.
Qed.

Examples

The propositions generated by verification_conditions are fairly big, and they contain many conjuncts that are essentially trivial.

(* Eval simpl in (verification_conditions (fun st => True) dec_while). *)
(* ====>
    (((fun st : state => True /\ bassn (BNot (BEq (AId X) (ANum 0))) st) ~~>
      (fun _ : state => True)) /\
     ((fun _ : state => True) ~~> (fun _ : state => True)) /\
     ((fun _ : state => True) ~~>
      assn_sub X (AMinus (AId X) (ANum 1)) (fun _ : state => True)) /\
     (fun st : state => True /\ ~ bassn (BNot (BEq (AId X) (ANum 0))) st) ~~>
     (fun st : state => True /\ ~ bassn (BNot (BEq (AId X) (ANum 0))) st)) /\
    (fun st : state => True /\ ~ bassn (BNot (BEq (AId X) (ANum 0))) st) ~~>
    (fun st : state => asnat (st X) = 0)
*)


We can certainly work with them using just the tactics we have so far, but we can make things much smoother with a bit of automation. We first define a custom verify tactic that applies split repeatedly to turn all the conjunctions into separate subgoals and then uses omega and eauto (a handy general-purpose automation tactic that we'll discuss in detail later) to deal with as many of them as possible.

Tactic Notation "verify" :=
  try apply verification_correct;
  repeat split;
  simpl; unfold assert_implies;
  unfold bassn in *; unfold beval in *; unfold aeval in *;
  unfold assn_sub; simpl in *;
  intros;
  repeat match goal with [H : _ /\ _ |- _] => destruct H end;
  try eauto; try omega.

What's left after verify does its thing is "just the interesting parts" of checking that the decorations are correct. In many cases, though, even these interesting parts can be dealt with fairly easily...

Theorem dec_while_correct :
  {{ fun st => True }}
  extract dec_while
  {{ fun st => asnat (st X) = 0 }}.
Proof.
  verify.
  destruct (asnat (st X)); auto. contradict H0. reflexivity.
Qed.

Don't expect yourself to understand the details of this proof script at this point. The purpose of showing it is just to give an impression of how proofs involving decorated programs actually look in practice.
Here's another:

Example subtract_slowly_dec (x:nat) (z:nat) : dcom := (
  WHILE BNot (BEq (AId X) (ANum 0))
  DO
                   {{ fun st => asnat (st Z) - asnat (st X) = z - x
                                /\ bassn (BNot (BEq (AId X) (ANum 0))) st }}
                   =>
                   {{ fun st => (asnat (st Z) - 1)
                                - (asnat (st X) - 1)
                                = z - x }}
     Z ::= AMinus (AId Z) (ANum 1)
                   {{ fun st => asnat (st Z)
                                - (asnat (st X) - 1)
                                = z - x }} ;
     X ::= AMinus (AId X) (ANum 1)
                   {{ fun st => asnat (st Z)
                                - asnat (st X)
                                = z - x }}
  END
(*                   {{ fun st =>   asnat (st Z) 
                                - asnat (st X) 
                                = z - x 
                                /\ ~ bassn (BNot (BEq (AId X) (ANum 0))) st }}
*)
=>
                   {{ fun st => asnat (st Z) = z - x }}
) % dcom.

Theorem subtract_slowly_dec_correct : forall x z,
  {{ fun st => asnat (st X) = x /\ asnat (st Z) = z }}
  extract (subtract_slowly_dec x z)
  {{ fun st => asnat (st Z) = z - x }}.
Proof.
  intros. verify.
    rewrite <- H.
    destruct (asnat (st X)). inversion H0. omega.
    destruct (asnat (st X)).
      omega.
      contradict H0. reflexivity.
Qed.

Finally, for a bigger example we can redo the proof of list_member_correct from above using our new tools.
Notice that the verify tactic leaves subgoals for each use of hoare_consequence -- that is, for each => that occurs in the decorated program. Each of these implications relies on a fact about lists, for example that l ++ [] = l. In other words, the Hoare logic infrastructure has taken care of the boilerplate reasoning about the execution of imperative programs, while the user has to prove lemmas that are specific to the problem domain (e.g. lists or numbers).

Definition list_member_decorated (n : nat) (l : list nat) : dcom := (
  =>
                   {{ fun st =>
                       st Y = VNat n
                        /\ exists p, p ++ aslist (st X) = l
                             /\ (st Z = VNat 1 <-> appears_in n p) }}
  WHILE BIsCons (AId X)
  DO
                   {{ fun st =>
                       (st Y = VNat n
                         /\ (exists p, p ++ aslist (st X) = l
                              /\ (st Z = VNat 1 <-> appears_in n p)))
                       /\ bassn (BIsCons (AId X)) st }}
    IFB (BEq (AId Y) (AHead (AId X))) THEN
                    {{ fun st =>
                      ((st Y = VNat n
                        /\ (exists p, p ++ aslist (st X) = l
                            /\ (st Z = VNat 1 <-> appears_in n p)))
                      /\ bassn (BIsCons (AId X)) st)
                      /\ bassn (BEq (AId Y) (AHead (AId X))) st }}
                    =>
                    {{ fun st =>
                        st Y = VNat n
                        /\ (exists p, p ++ tail (aslist (st X)) = l
                             /\ (VNat 1 = VNat 1 <-> appears_in n p)) }}
      Z ::= ANum 1
                     {{ fun st => st Y = VNat n
                          /\ (exists p, p ++ tail (aslist (st X)) = l
                               /\ (st Z = VNat 1 <-> appears_in n p)) }}
   ELSE
                    {{ fun st =>
                      ((st Y = VNat n
                        /\ (exists p, p ++ aslist (st X) = l
                              /\ (st Z = VNat 1 <-> appears_in n p)))
                      /\ bassn (BIsCons (AId X)) st)
                      /\ ~bassn (BEq (AId Y) (AHead (AId X))) st }}
                    =>
                    {{ fun st =>
                      st Y = VNat n
                      /\ (exists p, p ++ tail (aslist (st X)) = l
                           /\ (st Z = VNat 1 <-> appears_in n p)) }}
      SKIP
                    {{ fun st => st Y = VNat n
                        /\ (exists p, p ++ tail (aslist (st X)) = l
                             /\ (st Z = VNat 1 <-> appears_in n p)) }}
   FI ;
   X ::= (ATail (AId X))
                    {{ fun st =>
                        st Y = VNat n /\
                        (exists p : list nat, p ++ aslist (st X) = l
                          /\ (st Z = VNat 1 <-> appears_in n p)) }}
  END
(*                    {{ fun st =>
                      (st Y = VNat n 
                        /\ (exists p, p ++ aslist (st X) = l 
                             /\ (st Z = VNat 1 <-> appears_in n p)))
                      /\ ~bassn (BIsCons (AId X)) st }}
*)
=>
                    {{ fun st => st Z = VNat 1 <-> appears_in n l }}
) %dcom.

Theorem list_member_correct' : forall n l,
  {{ fun st => st X = VList l /\ st Y = VNat n /\ st Z = VNat 0 }}
  list_member
  {{ fun st => st Z = VNat 1 <-> appears_in n l }}.
Proof.
  intros n l.
  assert (list_member = extract (list_member_decorated n l))
    as Heq by reflexivity.
  rewrite Heq. clear Heq.
  verify.
  Case "The loop precondition holds.".
    exists []. simpl. split.
      rewrite H. reflexivity.
      rewrite H1. split; inversion 1.
  Case "IF taken".
     destruct H2 as [p [H3 H4]].
     (* We know X is non-nil. *)
     remember (aslist (st X)) as x.
     destruct x as [|h x'].
       inversion H1.
       exists (snoc p h).
       simpl. split.
          rewrite snoc_equation. assumption.
          split.
            rewrite H in H0.
            simpl in H0.
            rewrite (beq_true__eq _ _ H0).
            intros. apply appears_in_snoc1.
            intros. reflexivity.
  Case "If not taken".
     destruct H2 as [p [H3 H4]].
     (* We know X is non-nil. *)
     remember (aslist (st X)) as x.
     destruct x as [|h x'].
       inversion H1.
       exists (snoc p h).
       split.
         rewrite snoc_equation. assumption.
         split.
           intros. apply appears_in_snoc2. apply H4. assumption.
           intros Hai. destruct (appears_in_snoc3 _ _ _ Hai).
           SCase "later". apply H4. assumption.
           SCase "here (absurd)".
             subst.
             simpl in H0. rewrite H in H0. rewrite beq_nat_refl in H0.
             contradict H0. reflexivity.
  Case "The loop postcondition implies the desired conclusion (-> direction)".
     destruct H2 as [p [H3 H4]].
     unfold bassn in H1. simpl in H1.
     destruct (aslist (st X)) as [|h x'].
        rewrite append_nil in H3. subst. apply H4. assumption.
        contradict H1. reflexivity.
  Case "The loop postcondition implies the desired conclusion (<- direction)".
     destruct H2 as [p [H3 H4]].
     unfold bassn in H1. simpl in H1.
     destruct (aslist (st X)) as [|h x'].
        rewrite append_nil in H3. subst. apply H4. assumption.
        contradict H1. reflexivity.
Qed.