Library Exceptions

Calculation for arithmetic + exceptions.

Require Import List.
Require Import Tactics.

Syntax


Inductive Expr : Set :=
| Val : natExpr
| Add : ExprExprExpr
| Throw : Expr
| Catch : ExprExprExpr.

Semantics


Fixpoint eval (e: Expr) : option nat :=
  match e with
    | Val nSome n
    | Add x ymatch eval x with
                   | Some nmatch eval y with
                                 | Some mSome (n + m)
                                 | NoneNone
                               end
                   | NoneNone
                 end
    | ThrowNone
    | Catch x ymatch eval x with
                     | Some nSome n
                     | Noneeval y
                   end
  end.

Compiler


Inductive Code : Set :=
| PUSH : natCodeCode
| ADD : CodeCode
| FAIL : Code
| UNMARK : CodeCode
| MARK : CodeCodeCode
| HALT : Code.

Fixpoint comp' (e : Expr) (c : Code) : Code :=
  match e with
    | Val nPUSH n c
    | Add x ycomp' x (comp' y (ADD c))
    | ThrowFAIL
    | Catch x hMARK (comp' h c) (comp' x (UNMARK c))
  end.

Definition comp (e : Expr) : Code := comp' e HALT.

Virtual Machine


Inductive Elem : Set :=
| VAL : natElem
| HAN : CodeElem
.
Definition Stack : Set := list Elem.

Inductive Conf : Set :=
| conf : CodeStackConf
| fail : StackConf.

Notation "⟨ x , y ⟩" := (conf x y).
Notation "⟪ x ⟫" := (fail x ).

Reserved Notation "x ==> y" (at level 80, no associativity).
Inductive VM : ConfConfProp :=
| vm_push n c s : PUSH n c, s ==> c , VAL n :: s
| vm_add c s m n : ADD c, VAL m :: VAL n :: s ==> c, VAL (n + m) :: s
| vm_fail_val n s : VAL n :: s ==> s
| vm_fail s : FAIL, s ==> s
| vm_fail_han c s : HAN c :: s ==> c, s
| vm_unmark c n h s : UNMARK c, VAL n :: HAN h :: s ==> c, VAL n :: s
| vm_mark c h s : MARK h c, s ==> c, HAN h :: s
where "x ==> y" := (VM x y).

Hint Constructors VM.

Calculation

Boilerplate to import calculation tactics

Module VM <: Preorder.
Definition Conf := Conf.
Definition VM := VM.
End VM.
Module VMCalc := Calculation VM.
Import VMCalc.

Specification of the compiler

Theorem spec e c s : comp' e c, s
                       =>> match eval e with
                            | Some nc , VAL n :: s
                            | None s
                           end.

Setup the induction proof

Proof.
  intros.
  generalize dependent c.
  generalize dependent s.
  induction e;intros.

Calculation of the compiler

  begin
  c, VAL n :: s.
  <== { apply vm_push }
  PUSH n c, s.
  [].


  begin
   (match eval e1 with
     | Some mmatch eval e2 with
                  | Some n c, VAL (m + n) :: s
                  | None s
                  end
     | None s
     end).
  <<= { apply vm_add }
   (match eval e1 with
     | Some mmatch eval e2 with
                  | Some n ADD c, VAL n :: VAL m :: s
                  | None s
                  end
     | None s
     end).
  <<= { apply vm_fail_val }
   (match eval e1 with
     | Some mmatch eval e2 with
                  | Some n ADD c, VAL n :: VAL m :: s
                  | None VAL m :: s
                  end
     | None s
     end).
  <<= { apply IHe2 }
   (match eval e1 with
     | Some m comp' e2 (ADD c), VAL m :: s
     | None s
     end).
  <<= { apply IHe1 }
       comp' e1 (comp' e2 (ADD c)), s .
  [].


  begin
    s.
  <== { apply vm_fail }
     FAIL, s.
  [].


  begin
    (match eval e1 with
         | Some m c, VAL m :: s
         | Nonematch eval e2 with
                     | Some nc, VAL n :: s
                     | Nones
                   end
    end).
   <<= { apply IHe2 }
    (match eval e1 with
         | Some m c, VAL m :: s
         | Nonecomp' e2 c, s
    end).
   <<= { apply vm_fail_han }
    (match eval e1 with
         | Some m c, VAL m :: s
         | None HAN (comp' e2 c) :: s
    end).
   <<= { apply vm_unmark }
    (match eval e1 with
         | Some m UNMARK c, VAL m :: HAN (comp' e2 c) :: s
         | None HAN (comp' e2 c) :: s
    end).
   <<= { apply IHe1 }
        comp' e1 (UNMARK c), HAN (comp' e2 c) :: s.
   <<= { apply vm_mark }
        MARK (comp' e2 c) (comp' e1 (UNMARK c)), s.
   [].

Qed.

Soundness

Since the VM is defined as a small step operational semantics, we have to prove that the VM is deterministic and does not get stuck in order to derive soundness from the above theorem.

Lemma determ_vm : determ VM.
  intros C C1 C2 V. induction V; intro V'; inversion V'; subst; reflexivity.
Qed.

Lemma term_vm x : ¬ ( C, match x with
                             | Some nHALT , VAL n :: nil
                             | Nonenil
                           end ==> C).
Proof.
  destruct x; intro Contra; destruct Contra; subst; inversion H.
Qed.

Theorem sound e C : comp e, nil =>>! CC = match eval e with
                                                  | Some nHALT , VAL n :: nil
                                                  | Nonenil
                                                end.
Proof.
  intros.
  pose (spec e HALT nil) as H'. unfold comp in ×. pose (determ_trc determ_vm) as D.
  unfold determ in D. eapply D. apply H. split. apply H'. apply term_vm.
Qed.