├── .gitignore ├── coq ├── _CoqProject ├── ListIndex.v ├── .coqdeps.d ├── Memory.v ├── Makefile.conf ├── LinearMemory.v ├── Arith.v ├── Tactics.v ├── Machine.v ├── ExceptionTwo.v ├── While.v ├── LambdaBad.v ├── ExceptionPartial.v ├── LambdaTruncate.v ├── Exception.v ├── Lambda.v ├── LocalState.v ├── WhileCoalesced.v ├── State.v ├── LambdaException.v └── Makefile ├── haskell ├── arith.lhs ├── except.lhs └── lambda.lhs └── README.md /.gitignore: -------------------------------------------------------------------------------- 1 | *.bak 2 | *.aux 3 | *.html 4 | *.css 5 | -------------------------------------------------------------------------------- /coq/_CoqProject: -------------------------------------------------------------------------------- 1 | -R . Comp 2 | Memory.v 3 | LinearMemory.v 4 | Machine.v 5 | Tactics.v 6 | ListIndex.v 7 | 8 | Arith.v 9 | Exception.v 10 | Lambda.v 11 | 12 | ExceptionPartial.v 13 | ExceptionTwo.v 14 | State.v 15 | LocalState.v 16 | LambdaTruncate.v 17 | LambdaBad.v 18 | LambdaException.v 19 | While.v 20 | WhileCoalesced.v 21 | -------------------------------------------------------------------------------- /coq/ListIndex.v: -------------------------------------------------------------------------------- 1 | Require Import List. 2 | 3 | Fixpoint nth {A} (l:list A) (i:nat) : option A := 4 | match l with 5 | | nil => None 6 | | x :: xs => match i with 7 | | 0 => Some x 8 | | S j => nth xs j 9 | end 10 | end. 11 | 12 | Lemma nth_map A B l i (f : A -> B) : nth (map f l) i = option_map f (nth l i). 13 | Proof. 14 | intros. 15 | generalize dependent i. 16 | induction l; intros; simpl. reflexivity. 17 | 18 | destruct i; simpl;auto. 19 | Qed. 20 | -------------------------------------------------------------------------------- /coq/.coqdeps.d: -------------------------------------------------------------------------------- 1 | Memory.vo Memory.glob Memory.v.beautified: Memory.v 2 | Memory.vio: Memory.v 3 | LinearMemory.vo LinearMemory.glob LinearMemory.v.beautified: LinearMemory.v Memory.vo 4 | LinearMemory.vio: LinearMemory.v Memory.vio 5 | Machine.vo Machine.glob Machine.v.beautified: Machine.v Memory.vo 6 | Machine.vio: Machine.v Memory.vio 7 | Tactics.vo Tactics.glob Tactics.v.beautified: Tactics.v Machine.vo Memory.vo 8 | Tactics.vio: Tactics.v Machine.vio Memory.vio 9 | ListIndex.vo ListIndex.glob ListIndex.v.beautified: ListIndex.v 10 | ListIndex.vio: ListIndex.v 11 | Arith.vo Arith.glob Arith.v.beautified: Arith.v Tactics.vo 12 | Arith.vio: Arith.v Tactics.vio 13 | Exception.vo Exception.glob Exception.v.beautified: Exception.v Tactics.vo 14 | Exception.vio: Exception.v Tactics.vio 15 | Lambda.vo Lambda.glob Lambda.v.beautified: Lambda.v ListIndex.vo Tactics.vo 16 | Lambda.vio: Lambda.v ListIndex.vio Tactics.vio 17 | ExceptionPartial.vo ExceptionPartial.glob ExceptionPartial.v.beautified: ExceptionPartial.v Tactics.vo 18 | ExceptionPartial.vio: ExceptionPartial.v Tactics.vio 19 | ExceptionTwo.vo ExceptionTwo.glob ExceptionTwo.v.beautified: ExceptionTwo.v Tactics.vo 20 | ExceptionTwo.vio: ExceptionTwo.v Tactics.vio 21 | State.vo State.glob State.v.beautified: State.v Tactics.vo 22 | State.vio: State.v Tactics.vio 23 | LocalState.vo LocalState.glob LocalState.v.beautified: LocalState.v Tactics.vo 24 | LocalState.vio: LocalState.v Tactics.vio 25 | LambdaTruncate.vo LambdaTruncate.glob LambdaTruncate.v.beautified: LambdaTruncate.v ListIndex.vo Tactics.vo 26 | LambdaTruncate.vio: LambdaTruncate.v ListIndex.vio Tactics.vio 27 | LambdaBad.vo LambdaBad.glob LambdaBad.v.beautified: LambdaBad.v ListIndex.vo Tactics.vo 28 | LambdaBad.vio: LambdaBad.v ListIndex.vio Tactics.vio 29 | LambdaException.vo LambdaException.glob LambdaException.v.beautified: LambdaException.v ListIndex.vo Tactics.vo 30 | LambdaException.vio: LambdaException.v ListIndex.vio Tactics.vio 31 | While.vo While.glob While.v.beautified: While.v ListIndex.vo Tactics.vo 32 | While.vio: While.v ListIndex.vio Tactics.vio 33 | WhileCoalesced.vo WhileCoalesced.glob WhileCoalesced.v.beautified: WhileCoalesced.v ListIndex.vo Tactics.vo 34 | WhileCoalesced.vio: WhileCoalesced.v ListIndex.vio Tactics.vio 35 | -------------------------------------------------------------------------------- /haskell/arith.lhs: -------------------------------------------------------------------------------- 1 | *** Haskell code for the arithmetic expressions example *** 2 | 3 | -------------------------------------------------------------------------- 4 | 5 | Source language: 6 | 7 | > data Expr = Val Int | Add Expr Expr 8 | > deriving Show 9 | 10 | -------------------------------------------------------------------------- 11 | 12 | Semantics: 13 | 14 | > eval :: Expr -> Int 15 | > eval (Val n) = n 16 | > eval (Add x y) = eval x + eval y 17 | 18 | -------------------------------------------------------------------------- 19 | 20 | Target language: 21 | 22 | > data Code = LOAD Int Code | STORE Reg Code | ADD Reg Code | HALT 23 | > deriving Show 24 | 25 | -------------------------------------------------------------------------- 26 | 27 | Compiler: 28 | 29 | > compile :: Expr -> Code 30 | > compile e = comp e first HALT 31 | > 32 | > comp :: Expr -> Reg -> Code -> Code 33 | > comp (Val n) r c = LOAD n c 34 | > comp (Add x y) r c = comp x r (STORE r (comp y (next r) (ADD r c))) 35 | 36 | -------------------------------------------------------------------------- 37 | 38 | Machine model: 39 | 40 | > type Mem = Reg -> Int 41 | > 42 | > type Reg = Int 43 | > 44 | > empty :: Mem 45 | > empty = \n -> 0 46 | > 47 | > set :: Reg -> Int -> Mem -> Mem 48 | > set r n m = \r' -> if r == r' then n else get r' m 49 | > 50 | > get :: Reg -> Mem -> Int 51 | > get r m = m r 52 | > 53 | > first :: Reg 54 | > first = 0 55 | > 56 | > next :: Reg -> Reg 57 | > next r = r+1 58 | 59 | -------------------------------------------------------------------------- 60 | 61 | Virtual machine: 62 | 63 | > type Conf = (Acc,Mem) 64 | > 65 | > type Acc = Int 66 | > 67 | > exec :: Code -> Conf -> Conf 68 | > exec (LOAD n c) (a,m) = exec c (n,m) 69 | > exec (STORE r c) (a,m) = exec c (a, set r a m) 70 | > exec (ADD r c) (a,m) = exec c (get r m + a, m) 71 | > exec HALT (a,m) = (a,m) 72 | 73 | -------------------------------------------------------------------------- 74 | 75 | Testing: 76 | 77 | > nine :: Expr 78 | > nine = Add (Add (Val 2) (Val 3)) (Val 4) 79 | > 80 | > nine' :: Expr 81 | > nine' = Add (Val 2) (Add (Val 3) (Val 4)) 82 | > 83 | > test :: Expr -> Acc 84 | > test e = fst $ exec (compile e) (0,empty) 85 | 86 | -------------------------------------------------------------------------- 87 | -------------------------------------------------------------------------------- /coq/Memory.v: -------------------------------------------------------------------------------- 1 | (* Abstract specification of memory *) 2 | 3 | Create HintDb memory. 4 | 5 | Module Type Memory. 6 | 7 | (* Types *) 8 | Parameter Reg : Set. 9 | Parameter Mem : Type -> Type. 10 | 11 | (* Operations *) 12 | Parameter empty : forall {T}, Mem T. 13 | Parameter set : forall {T}, Reg -> T -> Mem T -> Mem T. 14 | Parameter get : forall {T}, Reg -> Mem T -> option T. 15 | Parameter first : Reg. 16 | Parameter next : Reg -> Reg. 17 | 18 | (* Predicates *) 19 | Parameter freeFrom : forall {T}, Reg -> Mem T -> Prop. 20 | Parameter memle : forall {T}, Mem T -> Mem T -> Prop. 21 | 22 | (* Notations *) 23 | Declare Scope memory_scope. 24 | Notation "m ⊑ m'" := (memle m m') (at level 70) : memory_scope. 25 | Open Scope memory_scope. 26 | Notation "m ⊒ m'" := (m' ⊑ m) (at level 70) : memory_scope. 27 | Notation "m [ r ] = v" := (get r m = Some v) (at level 70). 28 | Notation "m [ r := v ]" := (set r v m) (at level 70). 29 | 30 | 31 | (* Property 1 *) 32 | Axiom empty_mem_free : forall T, freeFrom first (@empty T). 33 | 34 | (* Property 2 *) 35 | 36 | Axiom get_set : forall T (r : Reg) (v : T) (m : Mem T), 37 | get r (set r v m) = Some v. 38 | 39 | (* Property 3 *) 40 | 41 | Axiom memle_set : forall {T} (m : Mem T) r v, freeFrom r m -> m ⊑ set r v m. 42 | 43 | (* Property 4 *) 44 | 45 | Axiom freeFrom_set : forall {T} r (v : T) m, freeFrom r m -> freeFrom (next r) (set r v m). 46 | 47 | (* Property 5 *) 48 | 49 | Axiom memle_refl : forall {T} (m : Mem T), m ⊑ m. 50 | Axiom memle_trans : forall {T} (m m' u : Mem T), m ⊑ m' -> m' ⊑ u -> m ⊑ u. 51 | 52 | (* Property 6 *) 53 | 54 | Axiom set_monotone : forall {T} (m m' : Mem T) r v, m ⊑ m' -> set r v m ⊑ set r v m' . 55 | 56 | (* Property 7 *) 57 | 58 | Axiom memle_get : forall {T} (m m' : Mem T) r v, m ⊑ m' -> get r m = Some v -> get r m' = Some v. 59 | 60 | 61 | Hint Resolve memle_set memle_get memle_refl set_monotone memle_trans freeFrom_set empty_mem_free : memory. 62 | End Memory. 63 | 64 | 65 | Lemma rel_eq {T} {R : T -> T -> Prop} x y y' : R x y' -> y = y' -> R x y. 66 | Proof. intros. subst. auto. 67 | Qed. 68 | 69 | 70 | (* Additional axioms not used in the paper. *) 71 | Module Type MAxioms (Import mem: Memory). 72 | Axiom set_set : forall T (r : Reg) (v v' : T) (m : Mem T), 73 | set r v (set r v' m) = set r v m. 74 | 75 | Ltac apply_eq t := eapply rel_eq; [apply t | repeat rewrite set_set; auto]. 76 | End MAxioms. 77 | 78 | Module Type SetMem := Memory <+ MAxioms. 79 | 80 | (* Extends the memory type with a truncation operation for proper 81 | modelling of stack frames. *) 82 | Module Type Truncate (Import mem:Memory). 83 | Parameter truncate : forall {T}, Reg -> Mem T -> Mem T. 84 | 85 | Axiom truncate_monotone : forall {T} (m m' : Mem T) r, m ⊑ m' -> truncate r m ⊑ truncate r m'. 86 | Axiom truncate_set : forall {T} (m : Mem T) v r , freeFrom r m -> truncate r (set r v m) = m. 87 | 88 | Hint Resolve truncate_monotone truncate_set : memory. 89 | 90 | End Truncate. 91 | 92 | Module Type TruncMem := SetMem <+ Truncate. 93 | -------------------------------------------------------------------------------- /README.md: -------------------------------------------------------------------------------- 1 | # Calculating compilers for register machines 2 | 3 | This repository contains the supplementary material for the paper 4 | *Calculating Correct Compilers II: Return of the Register Machines* by 5 | Patrick Bahr and Graham Hutton. The material includes Coq 6 | formalisations of all calculations in the paper along with the full 7 | Haskell source code of the calculated compilers. In addition, we also 8 | include Coq formalisations for calculations that were mentioned but 9 | not explicitly carried out in the paper. 10 | 11 | ## Coq formalisation 12 | 13 | The Coq formalisation is located in the [coq](coq) subfolder. Below we 14 | list the Coq files for the calculations from the paper and the 15 | specification of the memory model: 16 | 17 | - [Memory.v](coq/Memory.v): the (axiomatic) memory model (Section 2) 18 | - [Arith.v](coq/Arith.v): arithmetic expressions (section 3) 19 | - [Exception.v](coq/Exception.v): arithmetic expressions + exceptions (section 4) 20 | - [Lambda.v](coq/Lambda.v): call-by-value lambda calculus (section 5) 21 | 22 | We also include compiler calculations for additional languages: 23 | 24 | - [State.v](coq/State.v): global state + exceptions 25 | - [LocalState.v](coq/LocalState.v): local state + exceptions 26 | - [LambdaBad.v](coq/LambdaBad.v): call-by-value lambda calculus 27 | without a stack which yields an unsatisfactory compiler and machine 28 | - [LambdaTruncate.v](coq/LambdaTruncate.v): variant of 29 | [Lambda.v](coq/Lambda.v) is more realistic as it only copies 30 | a bounded set of registers to the stack (via truncation) 31 | - [LambdaException.v](coq/LambdaException.v): lambda calculus + 32 | exceptions, based on [LambdaTruncate.v](coq/LambdaTruncate.v) 33 | - [While.v](coq/While.v): state + while loops 34 | - [WhileCoalesced.v](coq/WhileCoalesced.v): Variant of [While.v](coq/While.v) with a 35 | semantics better suited for calculation 36 | - [ExceptionPartial.v](coq/ExceptionPartial.v): variant of 37 | [Exception.v](coq/Exception.v) with an expclicit empty exception 38 | handler (for unhandled exceptions) 39 | - [ExceptionTwo.v](coq/ExceptionTwo.v): variant of 40 | [Exception.v](coq/Exception.v) that compiles exceptions away by 41 | using two code continuations 42 | 43 | The remaining files are used to define the Coq tactics to support 44 | reasoning in calculation style and to provide auxiliary concepts: 45 | 46 | - [Tactics.v](coq/Tactics.v): tactics for calculation style proofs 47 | - [Machine.v](coq/Machine.v): auxiliary definitions and tactics for 48 | virtual machines 49 | - [LinearMemory.v](coq/LinearMemory.v): instantiation of the memory 50 | model (thus proving its consistency) 51 | - [ListIndex.v](coq/ListIndex.v): definitions to index elements in a list 52 | 53 | ## Haskell code 54 | 55 | The [haskell](haskell) subfolder contains the source code of the three 56 | compilers calculated in the paper: 57 | 58 | - [arith.lhs](haskell/arith.lhs): arithmetic expressions (section 3) 59 | - [except.lhs](haskell/except.lhs): arithmetic expressions + exceptions (section 4) 60 | - [lambda.lhs](haskell/lambda.lhs): lambda calculus (section 5) 61 | -------------------------------------------------------------------------------- /coq/Makefile.conf: -------------------------------------------------------------------------------- 1 | # This configuration file was generated by running: 2 | # coq_makefile -f _CoqProject -o makefile 3 | 4 | 5 | ############################################################################### 6 | # # 7 | # Project files. # 8 | # # 9 | ############################################################################### 10 | 11 | COQMF_VFILES = Memory.v LinearMemory.v Machine.v Tactics.v ListIndex.v Arith.v Exception.v Lambda.v ExceptionPartial.v ExceptionTwo.v State.v LocalState.v LambdaTruncate.v LambdaBad.v LambdaException.v While.v WhileCoalesced.v 12 | COQMF_MLIFILES = 13 | COQMF_MLFILES = 14 | COQMF_ML4FILES = 15 | COQMF_MLPACKFILES = 16 | COQMF_MLLIBFILES = 17 | COQMF_CMDLINE_VFILES = 18 | 19 | ############################################################################### 20 | # # 21 | # Path directives (-I, -R, -Q). # 22 | # # 23 | ############################################################################### 24 | 25 | COQMF_OCAMLLIBS = 26 | COQMF_SRC_SUBDIRS = 27 | COQMF_COQLIBS = -R . Comp 28 | COQMF_COQLIBS_NOML = -R . Comp 29 | COQMF_CMDLINE_COQLIBS = 30 | 31 | ############################################################################### 32 | # # 33 | # Coq configuration. # 34 | # # 35 | ############################################################################### 36 | 37 | COQMF_LOCAL=0 38 | COQMF_COQLIB=/Users/paba/.opam/default/lib/coq/ 39 | COQMF_DOCDIR=/Users/paba/.opam/default/doc/ 40 | COQMF_OCAMLFIND=/Users/paba/.opam/default/bin/ocamlfind 41 | COQMF_CAMLP5O=/Users/paba/.opam/default/bin/camlp5o 42 | COQMF_CAMLP5BIN=/Users/paba/.opam/default/bin/ 43 | COQMF_CAMLP5LIB=/Users/paba/.opam/default/lib/camlp5 44 | COQMF_CAMLP5OPTIONS=-loc loc 45 | COQMF_CAMLFLAGS=-thread -rectypes -w +a-4-9-27-41-42-44-45-48-50-58-59 -safe-string -strict-sequence 46 | COQMF_HASNATDYNLINK=true 47 | COQMF_COQ_SRC_SUBDIRS=config dev lib clib kernel library engine pretyping interp parsing proofs tactics toplevel printing grammar ide stm vernac plugins/btauto plugins/cc plugins/derive plugins/extraction plugins/firstorder plugins/funind plugins/ltac plugins/micromega plugins/nsatz plugins/omega plugins/quote plugins/romega plugins/rtauto plugins/setoid_ring plugins/ssr plugins/ssrmatching plugins/syntax plugins/xml 48 | COQMF_WINDRIVE= 49 | 50 | ############################################################################### 51 | # # 52 | # Extra variables. # 53 | # # 54 | ############################################################################### 55 | 56 | COQMF_OTHERFLAGS = 57 | COQMF_INSTALLCOQDOCROOT = Comp 58 | -------------------------------------------------------------------------------- /haskell/except.lhs: -------------------------------------------------------------------------------- 1 | *** Haskell code for the exceptions example *** 2 | 3 | Note: renamed 'fail' to 'myfail' to avoid clashing with prelude 4 | 5 | -------------------------------------------------------------------------- 6 | 7 | Source language: 8 | 9 | > data Expr = Val Int | Add Expr Expr | Throw | Catch Expr Expr 10 | > deriving Show 11 | 12 | -------------------------------------------------------------------------- 13 | 14 | Semantics: 15 | 16 | > eval :: Expr -> Maybe Int 17 | > eval (Val n) = Just n 18 | > eval (Add x y) = case eval x of 19 | > Just n -> case eval y of 20 | > Just m -> Just (n + m) 21 | > Nothing -> Nothing 22 | > Nothing -> Nothing 23 | > eval Throw = Nothing 24 | > eval (Catch x y) = case eval x of 25 | > Just n -> Just n 26 | > Nothing -> eval y 27 | 28 | -------------------------------------------------------------------------- 29 | 30 | Target language: 31 | 32 | > data Code = HALT | LOAD Int Code | STORE Reg Code | ADD Reg Code | 33 | > THROW | MARK Reg Code Code | UNMARK Code 34 | > deriving Show 35 | 36 | -------------------------------------------------------------------------- 37 | 38 | Compiler: 39 | 40 | > compile :: Expr -> Code 41 | > compile e = comp e first HALT 42 | 43 | > comp :: Expr -> Reg -> Code -> Code 44 | > comp (Val n) r c = LOAD n c 45 | > comp (Add x y) r c = comp x r (STORE r (comp y (next r) (ADD r c))) 46 | > comp (Throw) r c = THROW 47 | > comp (Catch x y) r c = MARK r (comp y r c) (comp x (next r) (UNMARK c)) 48 | 49 | -------------------------------------------------------------------------- 50 | 51 | Machine model: 52 | 53 | > type Mem = Reg -> Val 54 | > 55 | > type Reg = Int 56 | > 57 | > data Val = VAL {val :: Int} | HAN {han :: Han} 58 | > 59 | > empty :: Mem 60 | > empty = \n -> VAL 0 61 | > 62 | > set :: Reg -> Val -> Mem -> Mem 63 | > set r v m = \r' -> if r == r' then v else get r' m 64 | > 65 | > get :: Reg -> Mem -> Val 66 | > get r m = m r 67 | > 68 | > first :: Reg 69 | > first = 0 70 | > 71 | > next :: Reg -> Reg 72 | > next r = r+1 73 | 74 | -------------------------------------------------------------------------- 75 | 76 | Virtual machine: 77 | 78 | > type Conf = (Acc,Han,Mem) 79 | > 80 | > type Acc = Int 81 | > 82 | > type Han = (Code,Reg) 83 | > 84 | > exec :: Code -> Conf -> Conf 85 | > exec HALT (a,h,m) = (a,h,m) 86 | > exec (LOAD n c) (a,h,m) = exec c (n,h,m) 87 | > exec (STORE r c) (a,h,m) = exec c (a,h, set r (VAL a) m) 88 | > exec (ADD r c) (a,h,m) = exec c (val (get r m) + a,h, m) 89 | > exec (THROW) (a,h,m) = myfail h m 90 | > exec (MARK r c' c) (a,h,m) = exec c (a,(c',r),set r (HAN h) m) 91 | > exec (UNMARK c) (a,(_,r),m) = exec c (a, han (get r m), m) 92 | > 93 | > myfail :: Han -> Mem -> Conf 94 | > myfail (c,r) m = exec c (0, han (get r m), m) 95 | 96 | -------------------------------------------------------------------------- 97 | 98 | Testing: 99 | 100 | > nine :: Expr 101 | > nine = Add (Add (Val 2) (Val 3)) (Val 4) 102 | > 103 | > nine' :: Expr 104 | > nine' = Add (Val 2) (Add (Val 3) (Val 4)) 105 | > 106 | > ok :: Expr 107 | > ok = Catch (Add (Val 1) Throw) (Val 2) 108 | > 109 | > crash :: Expr 110 | > crash = Catch (Add Throw (Val 1)) Throw 111 | > 112 | > test :: Expr -> Acc 113 | > test e = case exec (compile e) (0,(HALT,0),empty) of 114 | > (a,_,_) -> a 115 | 116 | -------------------------------------------------------------------------- 117 | -------------------------------------------------------------------------------- /coq/LinearMemory.v: -------------------------------------------------------------------------------- 1 | Require Import Memory. 2 | Require Import ZArith. 3 | Require Import FunctionalExtensionality. 4 | 5 | 6 | Module LinearMemory <: TruncMem. 7 | 8 | Definition Reg := nat. 9 | Definition Mem (T : Type) := nat -> option T. 10 | 11 | Definition first := 0. 12 | Definition next := S. 13 | Definition empty := fun T => (fun _ => None) : Mem T. 14 | Definition set {T} (r : Reg) (v : T) (m : Mem T) : Mem T := 15 | fun r' => if r' =? r then Some v else m r'. 16 | Definition get {T} (r : Reg) (m : Mem T) : option T := m r. 17 | 18 | 19 | 20 | Definition freeFrom {T} (r : Reg) (m : Mem T) := 21 | forall r', r <= r' -> m r' = None. 22 | 23 | Definition memle {T} (m m' : Mem T) := 24 | forall r v, m r = Some v -> m' r = Some v. 25 | 26 | Declare Scope memory_scope. 27 | 28 | Notation "s ⊑ t" := (memle s t) (at level 70) : memory_scope. 29 | Open Scope memory_scope. 30 | Notation "s ⊒ t" := (t ⊑ s) (at level 70) : memory_scope. 31 | 32 | (* Property 1 *) 33 | Lemma empty_mem_free: forall T, freeFrom first (empty T). 34 | Proof. 35 | unfold freeFrom, first, empty. intros. reflexivity. 36 | Qed. 37 | 38 | (* Property 2 *) 39 | 40 | Lemma get_set : forall T (r : Reg) (v : T) (s : Mem T), 41 | get r (set r v s) = Some v. 42 | Proof. 43 | unfold set, get. intros. rewrite <- beq_nat_refl. reflexivity. 44 | Qed. 45 | 46 | (* Property 3 *) 47 | 48 | Lemma memle_set : forall {T} (s : Mem T) r v, freeFrom r s -> s ⊑ set r v s. 49 | Proof. 50 | unfold freeFrom, set, memle. intros. 51 | remember (r0 =? r) as R. symmetry in HeqR. destruct R. apply beq_nat_true in HeqR. subst. 52 | rewrite H in H0;auto. inversion H0. assumption. 53 | Qed. 54 | 55 | (* Property 4 *) 56 | 57 | Lemma freeFrom_set : forall {T} r (v : T) s, freeFrom r s -> freeFrom (next r) (set r v s). 58 | Proof. 59 | unfold freeFrom, next, set. intros. 60 | remember (r' =? r) as R. symmetry in HeqR. destruct R. apply beq_nat_true in HeqR. subst. omega. 61 | apply H. omega. 62 | Qed. 63 | 64 | (* Property 5 *) 65 | 66 | Lemma memle_refl : forall {T} (s : Mem T), s ⊑ s. 67 | Proof. 68 | unfold memle. intros. assumption. 69 | Qed. 70 | 71 | Lemma memle_trans : forall {T} (s t u : Mem T), s ⊑ t -> t ⊑ u -> s ⊑ u. 72 | Proof. 73 | unfold memle. intros. eauto. 74 | Qed. 75 | 76 | (* Property 6 *) 77 | 78 | Lemma set_monotone : forall {T} (s t : Mem T) r v, s ⊑ t -> set r v s ⊑ set r v t . 79 | Proof. 80 | unfold set, memle. intros. destruct (r0 =? r); eauto. 81 | Qed. 82 | 83 | 84 | (* Property 7 *) 85 | 86 | Lemma memle_get : forall {T} (s t : Mem T) r v, s ⊑ t -> get r s = Some v -> get r t = Some v. 87 | Proof. 88 | unfold get, memle. intros. auto. 89 | Qed. 90 | 91 | 92 | (* Additional axiom *) 93 | 94 | Lemma set_set : forall T (r : Reg) (v v' : T) (s : Mem T), 95 | set r v (set r v' s) = set r v s. 96 | Proof. 97 | unfold set. intros. apply functional_extensionality. intros. 98 | destruct (x =? r); auto. 99 | Qed. 100 | 101 | (* truncation operation and axioms *) 102 | 103 | Definition truncate {T} (r : Reg) (m : Mem T) : Mem T := 104 | fun r' => if r <=? r' then None else m r'. 105 | 106 | Lemma truncate_monotone : forall {T} (s t : Mem T) r, s ⊑ t -> truncate r s ⊑ truncate r t. 107 | Proof. 108 | unfold truncate, memle. intros. 109 | destruct (r <=? r0);auto. 110 | Qed. 111 | 112 | Lemma truncate_set : forall {T} (s : Mem T) v r , freeFrom r s -> truncate r (set r v s) = s. 113 | Proof. 114 | unfold freeFrom, truncate, set. intros. 115 | apply functional_extensionality. intros. 116 | remember (r <=? x) as R. symmetry in HeqR. destruct R. apply leb_complete in HeqR. 117 | symmetry. auto. apply leb_complete_conv in HeqR. 118 | remember (x =? r) as R'. symmetry in HeqR'. destruct R'. apply beq_nat_true in HeqR'. subst. omega. reflexivity. 119 | Qed. 120 | 121 | End LinearMemory. 122 | 123 | 124 | -------------------------------------------------------------------------------- /coq/Arith.v: -------------------------------------------------------------------------------- 1 | (** Calculation of the simple arithmetic language. *) 2 | 3 | Require Import List. 4 | Require Import Tactics. 5 | Require Import Coq.Program.Equality. 6 | Module Arith (Import mem : Memory). 7 | 8 | 9 | 10 | (** * Syntax *) 11 | 12 | Inductive Expr : Set := 13 | | Val : nat -> Expr 14 | | Add : Expr -> Expr -> Expr. 15 | 16 | (** * Semantics *) 17 | 18 | Fixpoint eval (x: Expr) : nat := 19 | match x with 20 | | Val n => n 21 | | Add x y => eval x + eval y 22 | end. 23 | 24 | (** * Compiler *) 25 | 26 | Inductive Code : Set := 27 | | LOAD : nat -> Code -> Code 28 | | ADD : Reg -> Code -> Code 29 | | STORE : Reg -> Code -> Code 30 | | HALT : Code. 31 | 32 | Fixpoint comp (x : Expr) (r : Reg) (c : Code) : Code := 33 | match x with 34 | | Val n => LOAD n c 35 | | Add x y => comp x r (STORE r (comp y (next r) (ADD r c))) 36 | end. 37 | 38 | Definition compile (x : Expr) : Code := comp x first HALT. 39 | 40 | (** * Virtual Machine *) 41 | 42 | Inductive Conf : Type := conf : Code -> nat -> Mem nat -> Conf. 43 | 44 | Notation "⟨ x , y , z ⟩" := (conf x y z). 45 | 46 | Reserved Notation "x ==> y" (at level 80, no associativity). 47 | Inductive VM : Conf -> Conf -> Prop := 48 | | vm_load n a c m : 49 | ⟨LOAD n c, a, m⟩ ==> ⟨c, n, m⟩ 50 | | vm_add c m a r n : m[r] = n -> 51 | ⟨ADD r c, a, m⟩ ==> ⟨c, n + a, m⟩ 52 | | vm_store c m a r : 53 | ⟨STORE r c, a, m⟩ ==> ⟨c, a, m[r:=a]⟩ 54 | where "x ==> y" := (VM x y). 55 | 56 | 57 | Inductive cle : Conf -> Conf -> Prop := 58 | cle_mem x y z z' : z ⊑ z' -> cle ⟨ x , y , z ⟩ ⟨ x , y , z' ⟩. 59 | 60 | Hint Constructors cle : core. 61 | 62 | 63 | (** * Calculation *) 64 | 65 | (** Boilerplate to import calculation tactics *) 66 | Module VM <: Machine. 67 | Definition Conf := Conf. 68 | Definition Pre := cle. 69 | Definition Rel := VM. 70 | Lemma monotone : monotonicity cle VM. 71 | prove_monotonicity. Qed. 72 | Lemma preorder : is_preorder cle. 73 | prove_preorder. Qed. 74 | End VM. 75 | Module VMCalc := Calculation mem VM. 76 | Import VMCalc. 77 | 78 | 79 | (** Specification of the compiler comp *) 80 | 81 | Theorem spec : forall e r c a m, freeFrom r m -> ⟨comp e r c, a, m⟩ =|> ⟨c , eval e, m⟩ . 82 | 83 | 84 | (** Setup the induction proof *) 85 | 86 | Proof. 87 | induction e;intros. 88 | 89 | (** Calculation of the compiler *) 90 | 91 | (** - [x = Val n]: *) 92 | 93 | begin 94 | ⟨c , eval (Val n), m⟩. 95 | = {auto} 96 | ⟨ c, n, m⟩. 97 | <== { apply vm_load } 98 | ⟨ LOAD n c, a, m ⟩. 99 | []. 100 | 101 | (** - [x = Add e1 e2]: *) 102 | 103 | begin 104 | ⟨ c, eval (Add e1 e2), m ⟩. 105 | = {auto} 106 | ⟨ c, eval e1 + eval e2, m ⟩. 107 | ⊑ {auto} 108 | ⟨c, eval e1 + eval e2, m[r:=eval e1]⟩. 109 | <== {apply vm_add} 110 | ⟨ADD r c, eval e2, m[r:=eval e1]⟩. 111 | <|= {apply IHe2} 112 | ⟨comp e2 (next r) (ADD r c), eval e1, m[r:=eval e1]⟩. 113 | <== { apply vm_store} 114 | ⟨STORE r (comp e2 (next r) (ADD r c)), eval e1, m⟩. 115 | <|= { apply IHe1 } 116 | ⟨comp e1 r (STORE r (comp e2 (next r) (ADD r c))), a, m⟩. 117 | []. 118 | Qed. 119 | 120 | 121 | (** Specification of the whole compiler *) 122 | 123 | Theorem spec' e a : ⟨compile e, a, empty⟩ =|> ⟨HALT , eval e, empty⟩. 124 | Proof. 125 | begin 126 | ⟨ HALT, eval e, empty ⟩. 127 | <|= {apply spec; apply empty_mem_free} 128 | ⟨comp e first HALT, a, empty⟩. 129 | []. 130 | Qed. 131 | 132 | (** * Soundness *) 133 | 134 | (** Since the VM is defined as a small step operational semantics, we *) 135 | (* have to prove that the VM is deterministic and does not get stuck in *) 136 | (* order to derive soundness from the above theorem. *) 137 | 138 | 139 | Lemma determ_vm : determ VM. 140 | intros C C1 C2 V. induction V; intro V'; inversion V'; subst;eauto. 141 | rewrite H in *. inversion H5;subst. auto. 142 | Qed. 143 | 144 | 145 | Theorem sound x a C : ⟨compile x, a, empty⟩ =>>! C -> exists m, C = ⟨HALT, eval x, m⟩. 146 | Proof. 147 | intros M. 148 | pose (spec' x a). unfold Reach in *. repeat autodestruct. 149 | pose (determ_trc determ_vm) as D. 150 | unfold determ in D. inversion H0. subst. eexists. eapply D. apply M. split. 151 | apply H. intro Contra. destruct Contra. 152 | inversion H1. 153 | Qed. 154 | 155 | End Arith. 156 | -------------------------------------------------------------------------------- /haskell/lambda.lhs: -------------------------------------------------------------------------------- 1 | *** Haskell code for the lambda calculus example *** 2 | 3 | -------------------------------------------------------------------------- 4 | 5 | Source language: 6 | 7 | > data Expr = Val Int | Add Expr Expr | Var Int | Abs Expr | App Expr Expr 8 | > deriving Show 9 | 10 | -------------------------------------------------------------------------- 11 | 12 | Functional semantics: 13 | 14 | data Value = Num Int | Fun (Value -> Value) 15 | 16 | type Env = [Value] 17 | 18 | eval :: Expr -> Env -> Value 19 | eval (Val n) e = Num n 20 | eval (Add x y) e = case eval x e of 21 | Num n -> case eval y e of 22 | Num m -> Num (n+m) 23 | eval (Var i) e = e !! i 24 | eval (Abs x) e = Fun (\v -> eval x (v:e)) 25 | eval (App x y) e = case eval x e of 26 | Fun f -> f (eval y e) 27 | 28 | -------------------------------------------------------------------------- 29 | 30 | Relational semantics: 31 | 32 | > data Value = Num Int | Clo Expr Env 33 | > deriving Show 34 | > 35 | > type Env = [Value] 36 | > 37 | > eval :: Expr -> Env -> [Value] 38 | > eval (Val n) e = [Num n] 39 | > eval (Add x y) e = [Num (n+m) | Num n <- eval x e, Num m <- eval y e] 40 | > eval (Var i) e = [e !! i] 41 | > eval (Abs x) e = [Clo x e] 42 | > eval (App x y) e = [w | Clo x' e' <- eval x e, v <- eval y e, w <- eval x' (v:e')] 43 | 44 | -------------------------------------------------------------------------- 45 | 46 | Target language: 47 | 48 | > data Code = HALT | LOAD Int Code | LOOKUP Int Code 49 | > | STORE Reg Code | ADD Reg Code 50 | > | STC Reg Code | APP Reg Code 51 | > | ABS Code Code | RET 52 | > deriving Show 53 | 54 | -------------------------------------------------------------------------- 55 | 56 | Compiler: 57 | 58 | > compile :: Expr -> Code 59 | > compile e = comp e first HALT 60 | > 61 | > comp :: Expr -> Reg -> Code -> Code 62 | > comp (Val n) r c = LOAD n c 63 | > comp (Var i) r c = LOOKUP i c 64 | > comp (Add x y) r c = comp x r (STORE r (comp y (next r) (ADD r c))) 65 | > comp (App x y) r c = comp x r (STC r (comp y (next r) (APP r c))) 66 | > comp (Abs x) r c = ABS (comp x (next first) RET) c 67 | 68 | -------------------------------------------------------------------------- 69 | 70 | Machine model: 71 | 72 | > type Mem = Reg -> Val 73 | > 74 | > type Reg = Int 75 | > 76 | > data Val = NUM Int | CLO Code Env' 77 | > deriving Show 78 | > 79 | > empty :: Mem 80 | > empty = \n -> NUM 0 81 | > 82 | > set :: Reg -> Val -> Mem -> Mem 83 | > set r v m = \r' -> if r == r' then v else get r' m 84 | > 85 | > get :: Reg -> Mem -> Val 86 | > get r m = m r 87 | > 88 | > first :: Reg 89 | > first = 0 90 | > 91 | > next :: Reg -> Reg 92 | > next r = r+1 93 | 94 | -------------------------------------------------------------------------- 95 | 96 | Virtual machine: 97 | 98 | > type Conf = (Acc,Env',Lam,Mem) 99 | > 100 | > type Acc = Value' 101 | > 102 | > type Env' = [Value'] 103 | > 104 | > data Value' = Num' Int | Clo' Code Env' 105 | > deriving Show 106 | > 107 | > type Lam = [Mem] 108 | > 109 | > exec :: Code -> Conf -> Conf 110 | > exec HALT (a,e,l,m) = (a,e,l,m) 111 | > exec (LOAD n c) (a,e,l,m) = exec c (Num' n,e,l,m) 112 | > exec (LOOKUP i c) (a,e,l,m) = exec c (e !! i,e,l,m) 113 | > exec (STORE r c) (Num' n,e,l,m) = exec c (Num' n,e,l,set r (NUM n) m) 114 | > exec (ADD r c) (Num' a,e,l,m) = exec c (Num' (n+a),e,l,m) 115 | > where NUM n = get r m 116 | > exec (STC r c) (Clo' c' e',e,l,m) = exec c (Clo' c' e',e,l,set r (CLO c' e') m) 117 | > exec (APP r c) (a,e,l,m) = exec c' (a,a:e',m:l,set first (CLO c e) empty) 118 | > where CLO c' e' = get r m 119 | > exec (ABS c' c) (a,e,l,m) = exec c (Clo' c' e,e,l,m) 120 | > exec RET (a,e',m:l,m') = exec c (a,e,l,m) 121 | > where CLO c e = get first m' 122 | 123 | -------------------------------------------------------------------------- 124 | 125 | Conversion functions: 126 | 127 | > conv :: Value -> Value' 128 | > conv (Num n) = Num' n 129 | > conv (Clo x e) = Clo' (comp x (next first) RET) (convE e) 130 | > 131 | > convE :: Env -> Env' 132 | > convE e = map conv e 133 | 134 | -------------------------------------------------------------------------- 135 | 136 | Testing: 137 | 138 | > nine :: Expr 139 | > nine = Add (Val 2) (Add (Val 3) (Val 4)) 140 | > 141 | > nine' :: Expr 142 | > nine' = Add (Add (Val 2) (Val 3)) (Val 4) 143 | > 144 | > inc :: Expr 145 | > inc = Abs (Add (Var 0) (Val 1)) 146 | > 147 | > three :: Expr 148 | > three = App inc (Val 2) 149 | > 150 | > test :: Expr -> Acc 151 | > test e = case exec (compile e) (Num' 0,[],[],empty) of 152 | > (a,_,_,_) -> a 153 | 154 | -------------------------------------------------------------------------- 155 | -------------------------------------------------------------------------------- /coq/Tactics.v: -------------------------------------------------------------------------------- 1 | Definition Admit {A} : A. admit. Admitted. 2 | 3 | Ltac premise H := match type of H with 4 | forall (_ : ?A), _ => 5 | let P := fresh in assert A as P;[idtac | specialize (H P);clear P] 6 | end. 7 | 8 | 9 | Require Import List. 10 | 11 | Require Export Machine. 12 | Require Export Memory. 13 | Module Calculation (mem: Memory)(mod : Machine). 14 | Module Meta := MetaTheory mod. 15 | Export Meta. 16 | Export mem. 17 | 18 | 19 | 20 | Import ListNotations. 21 | 22 | Ltac autodestruct := match goal with 23 | | [ H : _ /\ _ |- _] => destruct H 24 | | [ H : exists _ , _ |- _] => destruct H 25 | end. 26 | Ltac rewr_assumption := idtac; match goal with 27 | | [R: _ = _ |- _ ] => first [rewrite R| rewrite <- R] 28 | end. 29 | 30 | 31 | 32 | 33 | Ltac eval_inv ev := let do_inv e H := (first [is_var e; fail 1|inversion H; subst; clear H]) 34 | in idtac; match goal with 35 | | [ H: ev ?e _ |- _ ] => do_inv e H 36 | | [ H: ev ?e _ _ |- _ ] => do_inv e H 37 | | [ H: ev ?e _ _ _ |- _ ] => do_inv e H 38 | | [ H: ev ?e _ _ _ _ |- _ ] => do_inv e H 39 | | _ => eauto 40 | end. 41 | Ltac smart_destruct x := first[is_var x;destruct x| let x' := fresh in remember x as x'; destruct x' ]. 42 | 43 | 44 | Ltac dist' ev := 45 | simpl in *; intros; subst; ev; 46 | match goal with 47 | | [ H : Some _ = Some _ |- _] => inversion H; clear H; dist' ev 48 | | [ H : Some _ = None |- _] => inversion H; dist' ev 49 | | [ H : None = Some _ |- _] => inversion H; dist' ev 50 | | [ H : pair _ _ = pair _ _ |- _] => inversion H; clear H; dist' ev 51 | | [ H: and _ _ |- _ ] => destruct H; dist' ev 52 | | [ H: ex _ |- _ ] => destruct H; dist' ev 53 | | [ H: or _ _ |- _ ] => destruct H; dist' ev 54 | | [ H: eq _ _ |- _ ] => rewrite H in *; dist' ev 55 | (* | [ |- and _ _ ] => split; repeat dist' ev *) 56 | (* | [ |- _ <-> _ ] => split; dist' ev *) 57 | (* | [ |- ex _ ] => eexists; dist' ev *) 58 | (* | [ |- or _ _] => solve [right;dist' ev|left; dist' ev] *) 59 | | [ |- context [let _ := ?x in _] ] => smart_destruct x;dist' ev 60 | | [ |- context [match ?x with _ => _ end]] => smart_destruct x; dist' ev 61 | | _ => idtac 62 | end. 63 | 64 | Ltac dist := dist' eauto. 65 | 66 | 67 | Ltac solve_memle t := 68 | solve[ 69 | apply memle_set; 70 | match goal with 71 | | [H: freeFrom _ _ |- _] => apply H; t 72 | (* | [H: exists v, empty_mem _ [_] = v |- _ ] => apply empty_mem_free in H; contradiction; t *) 73 | | _ => t 74 | end 75 | | t 76 | ]. 77 | 78 | 79 | Ltac check_exp x y := let h := fresh "check" in assert (h: x = y) by dist; clear h. 80 | 81 | Ltac check_rel Bidir Rel := first [check_exp Bidir Rel| 82 | fail 2 "wrong goal; expected relation =>> but found" Rel]. 83 | 84 | Tactic Notation "[]" := apply Reach_refl. 85 | 86 | 87 | Tactic Notation (at level 2) "⊑" "{?}" constr(e2) := 88 | match goal with 89 | | [|- ?Rel ?lhs ?rhs] => check_rel Reach Rel; 90 | let h := fresh "rewriting" in 91 | assert(Pre rhs e2) 92 | | _ => fail 1 "goal is not a VM" 93 | end. 94 | 95 | Tactic Notation (at level 2) "=" "{?}" constr(e2) := 96 | match goal with 97 | | [|- ?Rel ?lhs ?rhs] => check_rel Reach Rel; 98 | let h := fresh "rewriting" in 99 | assert(rhs = e2) 100 | | _ => fail 1 "goal is not a VM" 101 | end. 102 | 103 | Tactic Notation (at level 2) "<==" "{?}" constr(e2) := 104 | match goal with 105 | | [|- ?Rel ?lhs ?rhs] => check_rel Reach Rel; 106 | let h := fresh "rewriting" in 107 | assert(e2 ==> rhs) 108 | | _ => fail 1 "goal is not a VM" 109 | end. 110 | 111 | 112 | 113 | Tactic Notation (at level 2) "<|=" "{?}" constr(e2) := 114 | match goal with 115 | | [|- ?Rel ?lhs ?rhs] => check_rel Reach Rel; 116 | first [let h := fresh "rewriting" in 117 | assert(h : Reach e2 rhs) | fail 2] 118 | | _ => fail 1 "goal is not a VM" 119 | end. 120 | 121 | 122 | 123 | Tactic Notation (at level 2) "<|=" "{{"tactic(t1) "}}" constr(e2) := 124 | match goal with 125 | | [|- ?Rel ?lhs ?rhs] => check_rel Reach Rel; 126 | first [let h := fresh "rewriting" in 127 | assert(h : Reach e2 rhs) by t1; 128 | apply (fun x => Reach_trans _ _ _ x h); clear h | fail 2] 129 | | _ => fail 1 "goal is not a VM" 130 | end. 131 | 132 | 133 | 134 | Tactic Notation (at level 2) "<|=" "{"tactic(t) "}" constr(e) := 135 | let t' := try solve [t;eauto with memory|apply Reach_refl;eauto] 136 | in 137 | <|= {{ dist' t' }} e . 138 | 139 | Tactic Notation (at level 2) "⊑" "{"tactic(t) "}" constr(e) := 140 | <|= {{ apply Reach_cle; dist; constructor; solve_memle t }} e . 141 | 142 | Tactic Notation (at level 2) "=" "{"tactic(t) "}" constr(e) := 143 | <|= {{ apply Reach_eq; dist' t}} e. 144 | 145 | 146 | Tactic Notation (at level 2) "<==" "{" tactic(t) "}" constr(e) := 147 | let tt := try solve[apply trc_step; t; eauto using get_set|apply trc_refl;eauto] 148 | in <|= {{ apply Reach_trc; dist' tt}} e. 149 | 150 | 151 | Tactic Notation (at level 2) "begin" constr(rhs) := 152 | match goal with 153 | | [|- ?Rel ?lhs ?rhs'] => check_rel Reach Rel; check_exp rhs rhs' 154 | | _ => fail 1 "rhs does not match" 155 | end. 156 | End Calculation. 157 | -------------------------------------------------------------------------------- /coq/Machine.v: -------------------------------------------------------------------------------- 1 | Require Import Coq.Program.Equality. 2 | Require Export Memory. 3 | 4 | Definition is_preorder {Conf} (pre : Conf -> Conf -> Prop) : Prop 5 | := (forall c, pre c c) /\ (forall c1 c2 c3, pre c1 c2 -> pre c2 c3 -> pre c1 c3). 6 | 7 | 8 | Definition monotonicity {Conf} 9 | (pre : Conf -> Conf -> Prop) (vm : Conf -> Conf -> Prop) : 10 | Prop := forall (C1 C1' C2 : Conf), 11 | pre C1 C1' -> 12 | vm C1 C2 -> 13 | exists C2', vm C1' C2' /\ pre C2 C2'. 14 | 15 | Ltac prove_preorder := 16 | split;[ 17 | intros c; destruct c; 18 | repeat 19 | (match goal with 20 | | [H : _ * _ |- _] => destruct H 21 | end); 22 | eauto with memory 23 | | intros c1 c2 c3 L1 L2; destruct c1, c2, c3; 24 | inversion L1; inversion L2; subst; 25 | repeat 26 | (match goal with 27 | | [H : (_, _) = (_, _) |- _] => inversion H; clear H 28 | end); subst; 29 | eauto with memory]. 30 | 31 | Ltac prove_monotonicity1 := 32 | do 3 intro; intros Hle Step; 33 | dependent destruction Step; inversion Hle. 34 | 35 | Ltac prove_monotonicity2 := subst; 36 | eexists; (split; [try solve [econstructor; eauto with memory]| eauto with memory]). 37 | 38 | Ltac prove_monotonicity := prove_monotonicity1; prove_monotonicity2. 39 | 40 | Module Type Machine. 41 | Parameter Conf : Type. 42 | Parameter Pre : Conf -> Conf -> Prop. 43 | Parameter Rel : Conf -> Conf -> Prop. 44 | Parameter preorder : is_preorder Pre. 45 | Parameter monotone : monotonicity Pre Rel. 46 | End Machine. 47 | 48 | Require Import List. 49 | Require Import Relations. 50 | 51 | Module MetaTheory (machine : Machine). 52 | Export machine. 53 | Import ListNotations. 54 | 55 | Declare Scope machine_scope. 56 | Infix "==>" := Rel(at level 80, no associativity) : machine_scope. 57 | 58 | Definition trc := clos_refl_trans Conf Rel. 59 | 60 | Infix "=>>" := trc (at level 80, no associativity) : machine_scope. 61 | 62 | Open Scope machine_scope. 63 | 64 | Lemma trc_refl c : c =>> c. 65 | Proof. apply rt_refl. Qed. 66 | 67 | Lemma trc_step c1 c2 : c1 ==> c2 -> c1 =>> c2. 68 | Proof. apply rt_step. Qed. 69 | 70 | Lemma trc_step_trans c1 c2 c3 : c1 =>> c2 -> c2 ==> c3 -> c1 =>> c3. 71 | Proof. intros. eapply rt_trans; eauto using rt_step. Qed. 72 | 73 | Lemma trc_step_trans' c1 c2 c3 : c1 ==> c2 -> c2 =>> c3 -> c1 =>> c3. 74 | Proof. intros. eapply rt_trans; eauto using rt_step. Qed. 75 | 76 | Lemma trc_trans c1 c2 c3 : c1 =>> c2 -> c2 =>> c3 -> c1 =>> c3. 77 | Proof. apply rt_trans. Qed. 78 | 79 | 80 | Hint Resolve trc_step trc_step_trans : core. 81 | Hint Immediate trc_refl : core. 82 | 83 | Lemma trc_ind' : 84 | forall P : Conf -> Conf -> Prop, 85 | (forall c : Conf, P c c) -> 86 | (forall c1 c2 c3 : Conf, c1 ==> c2 -> c2 =>> c3 -> P c2 c3 -> P c1 c3) -> 87 | forall c c0 : Conf, c =>> c0 -> P c c0. 88 | Proof. 89 | intros X Y Z c1 c2 S. unfold trc in S. rewrite -> clos_rt_rt1n_iff in S. 90 | induction S; eauto. rewrite <- clos_rt_rt1n_iff in S. eauto. 91 | Qed. 92 | 93 | Infix "⊑" := Pre (at level 70) : machine_scope. 94 | Notation "x ⊒ y" := (Pre y x) (at level 70) : machine_scope. 95 | 96 | Lemma cle_trans (C1 C2 C3 : Conf) : C1 ⊑ C2 -> C2 ⊑ C3 -> C1 ⊑ C3. 97 | Proof. 98 | intros. pose preorder as P. unfold is_preorder in *. destruct P. eauto. 99 | Qed. 100 | Lemma cle_refl (C : Conf) : C ⊑ C. 101 | Proof. 102 | intros. pose preorder as P. unfold is_preorder in *. destruct P. eauto. 103 | Qed. 104 | 105 | Hint Resolve cle_refl : core. 106 | 107 | Lemma monotone_step (C1 C1' C2 : Conf) : 108 | C1 ⊑ C1' -> 109 | C1 ==> C2 -> 110 | exists C2', C1' ==> C2' /\ C2 ⊑ C2' . 111 | Proof. 112 | intros. pose monotone as M. unfold monotonicity in M. eauto. 113 | Qed. 114 | 115 | Lemma monotone_machine (C1 C1' C2 : Conf) : 116 | C1 ⊑ C1' -> 117 | C1 =>> C2 -> 118 | exists C2', C1' =>> C2' /\ C2 ⊑ C2' . 119 | Proof. 120 | intros I M. generalize dependent C1'. dependent induction M using trc_ind';intros. 121 | - exists C1'. split; eauto. 122 | - eapply monotone_step in I; eauto. destruct I as [m2' HS]. destruct HS as [S Ic']. 123 | apply IHM in Ic'. destruct Ic'. destruct H0. eexists. split. eapply trc_step_trans'; eassumption. assumption. 124 | Qed. 125 | 126 | Definition Reach (C1 C2 : Conf) : Prop := exists C, C1 =>> C /\ C ⊒ C2. 127 | 128 | Infix "=|>" := Reach (at level 80, no associativity). 129 | 130 | Lemma Reach_refl C : C =|> C. 131 | Proof. 132 | exists C. split; auto. 133 | Qed. 134 | 135 | Hint Resolve Reach_refl : core. 136 | 137 | Lemma Reach_eq C1 C2 : C1 = C2 -> C1 =|> C2. 138 | Proof. 139 | intros. subst. auto. 140 | Qed. 141 | 142 | 143 | 144 | Lemma Reach_trans C1 C2 C3 : C1 =|> C2 -> C2 =|> C3 -> C1 =|> C3. 145 | Proof. 146 | intros L1 L2. 147 | destruct L1 as [C1' L1]. destruct L1 as [S1 M1]. 148 | destruct L2 as [C2' L2]. destruct L2 as [S2 M2]. 149 | eapply monotone_machine in S2;[|eassumption]. destruct S2 as [C3' G]. destruct G as [S2' M2']. 150 | eexists. split. 151 | - eapply trc_trans. apply S1. apply S2'. 152 | - eapply cle_trans;eassumption. 153 | Qed. 154 | 155 | Lemma Reach_cle C1 C2 : C1 ⊒ C2 -> C1 =|> C2. 156 | Proof. 157 | intros L. eexists. split. apply trc_refl. assumption. 158 | Qed. 159 | 160 | Lemma Reach_trc C1 C2 : C1 =>> C2 -> C1 =|> C2. 161 | Proof. 162 | intros L. eexists. split. eassumption. apply cle_refl. 163 | Qed. 164 | 165 | Lemma Reach_vm C1 C2 : C1 ==> C2 -> C1 =|> C2. 166 | Proof. 167 | intros L. apply Reach_trc. apply trc_step. assumption. 168 | Qed. 169 | 170 | 171 | Definition determ {A} (R : A -> A -> Prop) : Prop := forall C C1 C2, R C C1 -> R C C2 -> C1 = C2. 172 | 173 | Definition trc' C C' := C =>> C' /\ ~ exists C'', C' ==> C''. 174 | 175 | Notation "x =>>! y" := (trc' x y) (at level 80, no associativity). 176 | 177 | 178 | Lemma determ_factor C1 C2 C3 : determ Rel -> C1 ==> C2 -> C1 =>>! C3 -> C2 =>> C3. 179 | Proof. 180 | unfold determ. intros. destruct H1. 181 | destruct H1 using trc_ind'. 182 | - exfalso. apply H2. eexists. eassumption. 183 | - assert (c2 = C2). eapply H. apply H1. apply H0. subst. auto. 184 | Qed. 185 | 186 | 187 | Lemma determ_trc : determ Rel -> determ trc'. 188 | Proof. 189 | intros. unfold determ. intros. destruct H0. induction H0 using trc_ind'. 190 | - destruct H1. destruct H0 using trc_ind'. reflexivity. exfalso. apply H2. eexists. eassumption. 191 | - apply IHtrc. apply H2. split. eapply determ_factor; eassumption. destruct H1. assumption. 192 | Qed. 193 | End MetaTheory. 194 | 195 | -------------------------------------------------------------------------------- /coq/ExceptionTwo.v: -------------------------------------------------------------------------------- 1 | (** Calculation of the simple arithmetic language with exceptions 2 | using two code continuations. *) 3 | 4 | Require Import List. 5 | Require Import Tactics. 6 | Require Import Coq.Program.Equality. 7 | Module Exception (Import mem : Memory). 8 | 9 | 10 | (** * Syntax *) 11 | 12 | Inductive Expr : Set := 13 | | Val : nat -> Expr 14 | | Add : Expr -> Expr -> Expr 15 | | Throw : Expr 16 | | Catch : Expr -> Expr -> Expr. 17 | 18 | (** * Semantics *) 19 | 20 | Fixpoint eval (x: Expr) : option nat := 21 | match x with 22 | | Val n => Some n 23 | | Add x1 x2 => match eval x1 with 24 | | Some m => match eval x2 with 25 | | Some n => Some (m + n) 26 | | None => None 27 | end 28 | | None => None 29 | end 30 | | Throw => None 31 | | Catch x h => match eval x with 32 | | Some n => Some n 33 | | None => eval h 34 | end 35 | end. 36 | 37 | (** * Compiler *) 38 | 39 | Inductive Code : Set := 40 | | LOAD : nat -> Code -> Code 41 | | ADD : Reg -> Code -> Code 42 | | STORE : Reg -> Code -> Code 43 | | THROW : Code 44 | | UNMARK : Code -> Code 45 | | MARK : Reg -> Code -> Code -> Code 46 | | HALT : Code. 47 | 48 | Fixpoint comp' (x : Expr) (r : Reg) (c : Code) (f : Code) : Code := 49 | match x with 50 | | Val n => LOAD n c 51 | | Add e1 e2 => comp' e1 r (STORE r (comp' e2 (next r) (ADD r c) f)) f 52 | | Throw => LOAD 0 f 53 | | Catch e1 e2 => comp' e1 r c (comp' e2 r c f) 54 | end. 55 | 56 | Definition comp (x : Expr) : Code := comp' x first HALT HALT. 57 | 58 | 59 | Inductive RVal : Set := 60 | | VAL : nat -> RVal. 61 | 62 | (** * Virtual Machine *) 63 | 64 | Inductive Conf' : Type := 65 | | conf : Code -> nat -> Conf'. 66 | 67 | Definition Conf : Type := Conf' * Mem RVal. 68 | 69 | Notation "⟨ x , y , s ⟩" := (conf x y, s). 70 | 71 | 72 | Reserved Notation "x ==> y" (at level 80, no associativity). 73 | Inductive VM : Conf -> Conf -> Prop := 74 | | vm_load n a c s : ⟨LOAD n c, a , s⟩ ==> ⟨c , n, s⟩ 75 | | vm_add c s a r n : s[r]= VAL n -> ⟨ADD r c, a , s⟩ ==> ⟨c , n + a, s⟩ 76 | | vm_store c s a r : ⟨STORE r c, a, s⟩ ==> ⟨c , a, s[r:=VAL a]⟩ 77 | where "x ==> y" := (VM x y). 78 | 79 | Inductive cle : Conf -> Conf -> Prop := 80 | | cle_mem f s s' : s ⊑ s' -> cle (f, s) (f, s'). 81 | 82 | Hint Constructors cle : core. 83 | 84 | 85 | (** * Calculation *) 86 | 87 | (** Boilerplate to import calculation tactics *) 88 | 89 | Module VM <: Machine. 90 | Definition Conf := Conf. 91 | Definition Pre := cle. 92 | Definition Rel := VM. 93 | Lemma monotone : monotonicity cle VM. 94 | prove_monotonicity. Qed. 95 | Lemma preorder : is_preorder cle. 96 | prove_preorder. Qed. 97 | End VM. 98 | 99 | Module VMCalc := Calculation mem VM. 100 | Import VMCalc. 101 | 102 | 103 | (** Specification of the compiler *) 104 | 105 | Theorem spec : forall e r c a f s, 106 | freeFrom r s -> 107 | ⟨comp' e r c f, a, s⟩ 108 | =|> 109 | match eval e with 110 | | Some n => ⟨c , n, s⟩ 111 | | None => ⟨f , 0, s⟩ 112 | end. 113 | 114 | 115 | (** Setup the induction proof *) 116 | 117 | Proof. 118 | induction e;intros. 119 | 120 | (** Calculation of the compiler *) 121 | 122 | (** - [x = Val n]: *) 123 | 124 | begin 125 | match eval (Val n) with 126 | | Some n' => ⟨c , n', s⟩ 127 | | None => ⟨c , 0, s⟩ 128 | end. 129 | = {auto} 130 | ⟨ c, n, s⟩. 131 | <== {apply vm_load} 132 | ⟨ LOAD n c, a, s ⟩. 133 | []. 134 | 135 | (** - [x = Add x1 x2]: *) 136 | 137 | begin 138 | match eval (Add e1 e2) with 139 | | Some n => ⟨c , n, s⟩ 140 | | None => ⟨f , 0, s⟩ 141 | end. 142 | = { auto } 143 | match eval e1 with 144 | | Some n => match eval e2 with 145 | | Some n' => ⟨c , n + n' ,s⟩ 146 | | None => ⟨f , 0 ,s⟩ 147 | end 148 | | None => ⟨f , 0 ,s⟩ 149 | end. 150 | ⊑ { auto } 151 | match eval e1 with 152 | | Some n => match eval e2 with 153 | | Some n' => ⟨c , n + n' ,s[r:=VAL n]⟩ 154 | | None => ⟨f , 0 ,s⟩ 155 | end 156 | | None => ⟨f , 0 ,s⟩ 157 | end. 158 | <== { apply vm_add } 159 | match eval e1 with 160 | | Some n => match eval e2 with 161 | | Some n' => ⟨ADD r c , n', s[r:=VAL n]⟩ 162 | | None => ⟨f , 0 ,s⟩ 163 | end 164 | | None => ⟨f , 0 ,s⟩ 165 | end. 166 | ⊑ { auto } 167 | match eval e1 with 168 | | Some n => match eval e2 with 169 | | Some n' => ⟨ADD r c , n', s[r:=VAL n]⟩ 170 | | None => ⟨f , 0 ,s[r:=VAL n]⟩ 171 | end 172 | | None => ⟨f , 0 ,s⟩ 173 | end. 174 | <|= { apply IHe2 } 175 | match eval e1 with 176 | | Some n => ⟨comp' e2 (next r) (ADD r c) f , n, s[r:=VAL n]⟩ 177 | | None => ⟨f , 0 ,s⟩ 178 | end. 179 | <== { apply vm_store } 180 | match eval e1 with 181 | | Some n => ⟨STORE r (comp' e2 (next r) (ADD r c) f) , n , s⟩ 182 | | None => ⟨f , 0 ,s⟩ 183 | end. 184 | <|= { apply IHe1 } 185 | ⟨comp' e1 r (STORE r (comp' e2 (next r) (ADD r c) f)) f, a, s⟩. 186 | []. 187 | 188 | (** - [x = Throw]: *) 189 | 190 | begin 191 | match eval Throw with 192 | | Some n => ⟨c , n, s⟩ 193 | | None => ⟨f , 0 ,s⟩ 194 | end. 195 | = { auto } 196 | ⟨f , 0 ,s⟩. 197 | <== { apply vm_load } 198 | ⟨LOAD 0 f , a ,s⟩. 199 | []. 200 | 201 | (** - [x = Catch x1 x2]: *) 202 | 203 | begin 204 | match eval (Catch e1 e2) with 205 | | Some n => ⟨c , n ,s⟩ 206 | | None => ⟨f , 0 ,s⟩ 207 | end. 208 | = { auto } 209 | match eval e1 with 210 | | Some n => ⟨c , n ,s⟩ 211 | | None => match eval e2 with 212 | | Some n' => ⟨c , n' , s⟩ 213 | | None => ⟨f , 0 ,s⟩ 214 | end 215 | end. 216 | <|= { apply IHe2 } 217 | match eval e1 with 218 | | Some n => ⟨c , n, s⟩ 219 | | None => ⟨comp' e2 r c f , 0 ,s⟩ 220 | end. 221 | 222 | <|= {apply IHe1} 223 | ⟨comp' e1 r c (comp' e2 r c f) , a , s⟩. 224 | []. 225 | Qed. 226 | 227 | 228 | (** * Soundness *) 229 | 230 | (** Since the VM is defined as a small step operational semantics, we *) 231 | (* have to prove that the VM is deterministic and does not get stuck in *) 232 | (* order to derive soundness from the above theorem. *) 233 | 234 | 235 | Lemma determ_vm : determ VM. 236 | intros C C1 C2 V. induction V; intro V'; inversion V'; subst;eauto. 237 | rewrite H in *; dist. 238 | Qed. 239 | 240 | 241 | Theorem sound x a s n C : freeFrom first s -> eval x = Some n -> ⟨comp x, a, s⟩ =>>! C -> exists s', C = ⟨HALT, n,s'⟩. 242 | Proof. 243 | intros F E M. 244 | pose (spec x first HALT a HALT s F). unfold Reach in *. repeat autodestruct. 245 | pose (determ_trc determ_vm) as D. 246 | unfold determ in D. inversion H0. subst. eexists. eapply D. apply M. 247 | rewrite E in *. 248 | split. inversion H1. subst. 249 | 250 | apply H. intro Contra. destruct Contra. 251 | inversion H3. 252 | Qed. 253 | 254 | End Exception. 255 | -------------------------------------------------------------------------------- /coq/While.v: -------------------------------------------------------------------------------- 1 | (** Calculation for a language with state and while loops. *) 2 | 3 | Require Import List. 4 | Require Import ListIndex. 5 | Require Import Tactics. 6 | Require Import Coq.Program.Equality. 7 | Module While (Import mem : Memory). 8 | 9 | (** * Syntax *) 10 | 11 | Inductive Expr : Set := 12 | | Val : nat -> Expr 13 | | Add : Expr -> Expr -> Expr 14 | | Get : Expr 15 | | Put : Expr -> Expr -> Expr 16 | | While : Expr -> Expr -> Expr. 17 | 18 | Definition State := nat. 19 | Definition Value := nat. 20 | 21 | Reserved Notation "⟪ x , q ⟫ ⇓ ⟪ y , q' ⟫" (at level 80, no associativity). 22 | 23 | Inductive eval : Expr -> State -> Value -> State -> Prop := 24 | | eval_val n q : ⟪Val n, q⟫ ⇓ ⟪n, q⟫ 25 | | eval_add q q' q'' x y m n : ⟪x, q⟫ ⇓ ⟪m,q'⟫ -> ⟪y,q'⟫ ⇓ ⟪n, q''⟫ -> ⟪Add x y, q⟫ ⇓ ⟪m + n, q''⟫ 26 | | eval_get q : ⟪Get, q⟫ ⇓ ⟪q, q⟫ 27 | | eval_put q q' q'' v n x y : ⟪x, q⟫ ⇓ ⟪n,q'⟫ -> ⟪y, n⟫ ⇓ ⟪v, q''⟫ -> ⟪Put x y , q⟫ ⇓ ⟪v,q''⟫ 28 | | eval_step x y q n m q' q'' v q''': 29 | ⟪x, q⟫ ⇓ ⟪n,q'⟫ -> 30 | n <> 0 -> ⟪y,q'⟫ ⇓ ⟪m,q''⟫ -> ⟪While x y, q''⟫ ⇓ ⟪v,q'''⟫ -> 31 | ⟪While x y, q⟫ ⇓ ⟪v,q'''⟫ 32 | | eval_while_stop x y q n q' : 33 | ⟪x, q⟫ ⇓ ⟪n,q'⟫ -> 34 | n = 0 -> 35 | ⟪While x y, q⟫ ⇓ ⟪0,q'⟫ 36 | where "⟪ x , q ⟫ ⇓ ⟪ y , q' ⟫" := (eval x q y q'). 37 | 38 | (** * Compiler *) 39 | 40 | Inductive Code : Set := 41 | | LOAD : nat -> Code -> Code 42 | | ADD : Reg -> Code -> Code 43 | | STORE : Reg -> Code -> Code 44 | | GET : Code -> Code 45 | | JUMP : Reg -> Code 46 | | JMPZ : Code -> Code -> Code 47 | | LABEL : Reg -> Code -> Code 48 | | PUT : Code -> Code 49 | | HALT : Code. 50 | 51 | Fixpoint comp' (e : Expr) (r : Reg) (c : Code) : Code := 52 | match e with 53 | | Val n => LOAD n c 54 | | Add x y => comp' x r (STORE r (comp' y (next r) (ADD r c))) 55 | | While x y => LABEL r (comp' x (next r) (JMPZ c (comp' y (next r) (JUMP r)))) 56 | | Get => GET c 57 | | Put x y => comp' x r (PUT (comp' y r c)) 58 | end. 59 | 60 | Definition comp (e : Expr) : Code := comp' e first HALT. 61 | 62 | 63 | Inductive Elem : Set := 64 | | VAL : Value -> Elem 65 | | CODE : Code -> Elem 66 | . 67 | 68 | 69 | 70 | (** * Virtual Machine *) 71 | 72 | Inductive Conf : Type := 73 | | conf : Code -> Value -> State -> Mem Elem -> Conf. 74 | 75 | Notation "⟨ c , a , q , s ⟩" := (conf c a q s). 76 | 77 | Reserved Notation "x ==> y" (at level 80, no associativity). 78 | Inductive VM : Conf -> Conf -> Prop := 79 | | vm_load n c s a q : ⟨LOAD n c, a, q, s⟩ ==> ⟨c, n, q, s⟩ 80 | | vm_add c m n r s q : s[r] = VAL m -> ⟨ADD r c, n, q, s⟩ 81 | ==> ⟨c, m + n, q, s⟩ 82 | | vm_store c v r s q : ⟨STORE r c, v, q, s⟩ 83 | ==> ⟨c, v, q, s[r:= VAL v]⟩ 84 | | vm_jump n r s q c : s[r] = CODE c -> ⟨JUMP r, n, q, s⟩ 85 | ==> ⟨c, n, q, s⟩ 86 | | vm_jmpz_zero s q c c' : ⟨JMPZ c' c, 0, q, s⟩ 87 | ==> ⟨c', 0, q, s⟩ 88 | | vm_jmpz_nzero n s q c c' : n <> 0 -> 89 | ⟨JMPZ c' c, n, q, s⟩ ==> ⟨c, n, q, s⟩ 90 | | vm_label n r s q c : ⟨LABEL r c, n, q, s⟩ ==> ⟨c, n, q, s[r:=CODE (LABEL r c)]⟩ 91 | | vm_get c s a q : ⟨GET c, a, q, s⟩ ==> ⟨c, q, q, s⟩ 92 | | vm_put c s a q : ⟨PUT c, a, q, s⟩ ==> ⟨c, a, a, s⟩ 93 | where "x ==> y" := (VM x y). 94 | 95 | (** Conversion functions from semantics to VM *) 96 | 97 | Inductive cle : Conf -> Conf -> Prop := 98 | | cle_mem c a e s s' : s ⊑ s' -> cle ⟨ c , a , e , s ⟩ ⟨ c , a , e , s' ⟩. 99 | 100 | Hint Constructors cle : core. 101 | 102 | 103 | (** * Calculation *) 104 | 105 | (** Boilerplate to import calculation tactics *) 106 | Module VM <: Machine. 107 | Definition Conf := Conf. 108 | Definition Pre := cle. 109 | Definition Rel := VM. 110 | Lemma monotone : monotonicity cle VM. 111 | prove_monotonicity. Qed. 112 | Lemma preorder : is_preorder cle. 113 | prove_preorder. Qed. 114 | End VM. 115 | Module VMCalc := Calculation mem VM. 116 | Import VMCalc. 117 | 118 | 119 | 120 | (** Specification of the compiler *) 121 | 122 | Theorem spec p v q q' r c a s : 123 | freeFrom r s -> ⟪p, q⟫ ⇓ ⟪v, q'⟫ -> 124 | ⟨comp' p r c, a, q, s⟩ =|> ⟨c , v, q', s⟩. 125 | 126 | (** Setup the induction proof *) 127 | 128 | Proof. 129 | intros F E. 130 | generalize dependent c. 131 | generalize dependent a. 132 | generalize dependent s. 133 | generalize dependent r. 134 | dependent induction E;intros. 135 | 136 | (** Calculation of the compiler *) 137 | 138 | (** - [(Val n,q) ⇓ (n,q)]: *) 139 | 140 | begin 141 | ⟨c, n , q, s⟩. 142 | <== { apply vm_load } 143 | ⟨LOAD n c, a, q, s⟩. 144 | []. 145 | 146 | (** - [(Add x y,q) ⇓ (m + n,q')]: *) 147 | 148 | begin 149 | ⟨c, m + n, q'', s⟩. 150 | ⊑ { auto } 151 | ⟨c, m + n, q'', s[r:=VAL m]⟩ . 152 | <== { apply vm_add } 153 | ⟨ADD r c, n, q'', s[r:=VAL m]⟩ . 154 | <|= { apply IHE2 } 155 | ⟨comp' y (next r) (ADD r c), m, q', s[r:= VAL m]⟩ . 156 | <== { apply vm_store } 157 | ⟨STORE r (comp' y (next r) (ADD r c)), m, q', s⟩. 158 | <|= { apply IHE1 } 159 | ⟨comp' x r (STORE r (comp' y (next r) (ADD r c))), a, q, s⟩. 160 | []. 161 | 162 | (** - [(Get,q) ⇓ (q,q)]: *) 163 | 164 | begin 165 | ⟨ c, q, q, s ⟩. 166 | <== {apply vm_get} 167 | ⟨ GET c, a, q, s ⟩. 168 | []. 169 | 170 | (** - [(Put x y,q) ⇓ (v,q'')]: *) 171 | begin 172 | ⟨ c, v, q'', s ⟩. 173 | <|= { apply IHE2 } 174 | ⟨ comp' y r c, n, n, s ⟩. 175 | ⊑ {auto} 176 | ⟨ comp' y r c, n, n, s ⟩. 177 | <== { apply vm_put } 178 | ⟨ PUT (comp' y r c), n, q', s ⟩. 179 | <|= { apply IHE1 } 180 | ⟨ comp' x r (PUT (comp' y r c)), a, q, s ⟩. 181 | []. 182 | 183 | 184 | (** - [(While x y ,q) ⇓ (v,q''')]: *) 185 | begin 186 | ⟨ c, v, q''', s ⟩. 187 | <|= { eapply IHE3 } 188 | ⟨ comp' (While x y) r c, m, q'', s ⟩. 189 | ⊑ { auto } 190 | ⟨ comp' (While x y) r c, m, q'', s[r:=CODE (comp' (While x y) r c)] ⟩. 191 | <== { apply vm_jump } 192 | ⟨ JUMP r, m, q'', s[r:=CODE (comp' (While x y) r c)] ⟩. 193 | <|= { apply IHE2 } 194 | ⟨ comp' y (next r) (JUMP r), n, q', s[r:=CODE (comp' (While x y) r c)] ⟩. 195 | <== { apply vm_jmpz_nzero } 196 | ⟨ JMPZ c (comp' y (next r) (JUMP r)), n, q', s[r:=CODE (comp' (While x y) r c)] ⟩. 197 | <|= { apply IHE1 } 198 | ⟨ comp' x (next r) (JMPZ c (comp' y (next r) (JUMP r))), a, q, s[r:=CODE (comp' (While x y) r c)] ⟩. 199 | <== { apply vm_label } 200 | ⟨ LABEL r (comp' x (next r) (JMPZ c (comp' y (next r) (JUMP r)))), a, q, s ⟩. 201 | []. 202 | 203 | begin 204 | ⟨ c, 0, q', s ⟩. 205 | <== { apply vm_jmpz_zero } 206 | ⟨ JMPZ c (comp' y (next r) (JUMP r)) , 0, q', s ⟩. 207 | ⊑ { auto } 208 | ⟨ JMPZ c (comp' y (next r) (JUMP r)) , 0, q', s[r:=CODE (comp' (While x y) r c)] ⟩. 209 | <|= {eapply IHE} 210 | ⟨ comp' x (next r) (JMPZ c (comp' y (next r) (JUMP r))) , a, q, s[r:=CODE (comp' (While x y) r c)] ⟩. 211 | <== { apply vm_label } 212 | ⟨ LABEL r (comp' x (next r) (JMPZ c (comp' y (next r) (JUMP r)))), a, q, s ⟩. 213 | []. 214 | Qed. 215 | 216 | (** * Soundness *) 217 | 218 | Lemma determ_vm : determ VM. 219 | intros C C1 C2 V. induction V; intro V'; inversion V'; subst; congruence. 220 | Qed. 221 | 222 | 223 | Definition terminates (p : Expr) : Prop := exists n q, ⟪p, 0⟫ ⇓ ⟪n,q⟫. 224 | 225 | Theorem sound p a s C : freeFrom first s -> terminates p -> ⟨comp p, a, 0, s⟩ =>>! C -> 226 | exists n s' q, C = ⟨HALT , n, q, s'⟩ /\ ⟪p,0⟫ ⇓ ⟪n,q⟫. 227 | Proof. 228 | unfold terminates. intros F T M. destruct T as [n T]. destruct T as [q T]. 229 | pose (spec p n 0 q first HALT a s F T) as H'. 230 | unfold Reach in *. repeat autodestruct. 231 | pose (determ_trc determ_vm) as D. 232 | unfold determ in D. inversion H0. subst. exists n. eexists. exists q. split. eapply D. apply M. split. 233 | unfold comp. 234 | simpl in *. apply H. intro Contra. destruct Contra. 235 | inversion H1. assumption. 236 | Qed. 237 | 238 | End While. 239 | -------------------------------------------------------------------------------- /coq/LambdaBad.v: -------------------------------------------------------------------------------- 1 | (** Calculation for the lambda calculus + arithmetic without stack. *) 2 | 3 | Require Import List. 4 | Require Import ListIndex. 5 | Require Import Tactics. 6 | Require Import Coq.Program.Equality. 7 | Module Lambda (Import mem : SetMem). (* memory with additional axiom set_set *) 8 | 9 | (** * Syntax *) 10 | 11 | Inductive Expr : Set := 12 | | Val : nat -> Expr 13 | | Add : Expr -> Expr -> Expr 14 | | Var : nat -> Expr 15 | | Abs : Expr -> Expr 16 | | App : Expr -> Expr -> Expr. 17 | 18 | (** * Semantics *) 19 | 20 | (** The evaluator for this language is given as follows (as in the 21 | paper): 22 | << 23 | type Env = [Value] 24 | data Value = Num Int | Fun (Value -> Value) 25 | 26 | 27 | eval :: Expr -> Env -> Value 28 | eval (Val n) e = Num n 29 | eval (Add x y) e = case eval x e of 30 | Num n -> case eval y e of 31 | Num m -> Num (n+m) 32 | eval (Var i) e = e !! i 33 | eval (Abs x) e = Fun (\v -> eval x (v:e)) 34 | eval (App x y) e = case eval x e of 35 | Fun f -> f (eval y e) 36 | >> 37 | After defunctionalisation and translation into relational form we 38 | obtain the semantics below. *) 39 | 40 | Inductive Value : Set := 41 | | Num : nat -> Value 42 | | Clo : Expr -> list Value -> Value. 43 | 44 | Definition Env := list Value. 45 | 46 | Reserved Notation "x ⇓[ e ] y" (at level 80, no associativity). 47 | 48 | Inductive eval : Expr -> Env -> Value -> Prop := 49 | | eval_val e n : Val n ⇓[e] Num n 50 | | eval_add e x y m n : x ⇓[e] Num m -> y ⇓[e] Num n -> Add x y ⇓[e] Num (m + n) 51 | | eval_var e i v : nth e i = Some v -> Var i ⇓[e] v 52 | | eval_abs e x : Abs x ⇓[e] Clo x e 53 | | eval_app e e' x x' x'' y y' : x ⇓[e] Clo x' e' -> y ⇓[e] y' -> x' ⇓[y' :: e'] x'' -> App x y ⇓[e] x'' 54 | where "x ⇓[ e ] y" := (eval x e y). 55 | 56 | (** * Compiler *) 57 | 58 | Inductive Code : Set := 59 | | LOAD : nat -> Code -> Code 60 | | ADD : Reg -> Code -> Code 61 | | STORE : Reg -> Code -> Code 62 | | LOOKUP : nat -> Code -> Code 63 | | RET : Reg -> Code 64 | | CALL : Reg -> Code -> Code 65 | | FUN : (Reg -> Code) -> Code -> Code 66 | | HALT : Code. 67 | 68 | Fixpoint comp' (e : Expr) (r : Reg) (c : Code) : Code := 69 | match e with 70 | | Val n => LOAD n c 71 | | Add x y => comp' x r (STORE r (comp' y (next r) (ADD r c))) 72 | | Var i => LOOKUP i c 73 | | App x y => comp' x r (STORE r (comp' y (next r) (CALL r c))) 74 | | Abs x => FUN (fun r' => comp' x (next r') (RET r')) c 75 | end. 76 | 77 | Definition comp (e : Expr) : Code := comp' e first HALT. 78 | 79 | (** * Virtual Machine *) 80 | 81 | Inductive Value' : Set := 82 | | Num' : nat -> Value' 83 | | Clo' : (Reg -> Code) -> list Value' -> Value'. 84 | 85 | Definition Env' := list Value'. 86 | 87 | Inductive RVal : Set := 88 | | VAL : Value' -> RVal 89 | | CLO : Code -> Env' -> RVal 90 | . 91 | 92 | Inductive Conf : Type := 93 | | conf : Code -> Value' -> Env' -> Mem RVal -> Conf. 94 | 95 | Notation "⟨ c , a , e , s ⟩" := (conf c a e s). 96 | 97 | Reserved Notation "x ==> y" (at level 80, no associativity). 98 | Inductive VM : Conf -> Conf -> Prop := 99 | | vm_load n c s a e : ⟨LOAD n c, a, e, s⟩ ==> ⟨c, Num' n, e, s⟩ 100 | | vm_add c m n r s e : s[r] = VAL (Num' m) -> ⟨ADD r c, Num' n, e, s⟩ 101 | ==> ⟨c, Num'(m + n), e, s⟩ 102 | | vm_store c v r s e : ⟨STORE r c, v, e, s⟩ 103 | ==> ⟨c, v, e, s[r:=VAL v]⟩ 104 | | vm_lookup e i c v a s : nth e i = Some v -> ⟨LOOKUP i c, a, e, s⟩ ==> ⟨c, v, e, s⟩ 105 | | vm_ret a r c e e' s : s[r] = CLO c e -> ⟨RET r, a, e', s⟩ ==> ⟨c, a, e, s⟩ 106 | | vm_call c c' e e' v r s : 107 | s[r]=VAL (Clo' c' e') -> 108 | ⟨CALL r c, v, e,s⟩ ==> ⟨c' r, Num' 0, v :: e', s[r:=CLO c e]⟩ 109 | | vm_fun a c c' s e : ⟨FUN c' c, a, e, s⟩ ==> ⟨c, Clo' c' e, e, s⟩ 110 | where "x ==> y" := (VM x y). 111 | 112 | (** Conversion functions from semantics to VM *) 113 | 114 | Fixpoint conv (v : Value) : Value' := 115 | match v with 116 | | Num n => Num' n 117 | | Clo x e => Clo' (fun r => comp' x (next r) (RET r)) (map conv e) 118 | end. 119 | 120 | Definition convE : Env -> Env' := map conv. 121 | 122 | 123 | Inductive cle : Conf -> Conf -> Prop := 124 | | cle_mem c a e s s' : s ⊑ s' -> cle ⟨ c , a , e , s ⟩ ⟨ c , a , e , s' ⟩. 125 | 126 | Hint Constructors cle : core. 127 | 128 | 129 | (** * Calculation *) 130 | 131 | (** Boilerplate to import calculation tactics *) 132 | Module VM <: Machine. 133 | Definition Conf := Conf. 134 | Definition Pre := cle. 135 | Definition Rel := VM. 136 | 137 | Lemma monotone : monotonicity cle VM. 138 | prove_monotonicity. Qed. 139 | Lemma preorder : is_preorder cle. 140 | prove_preorder. Qed. 141 | End VM. 142 | Module VMCalc := Calculation mem VM. 143 | Import VMCalc. 144 | 145 | (** Specification of the compiler *) 146 | 147 | Theorem spec p v e r c a s : 148 | freeFrom r s -> p ⇓[e] v -> 149 | ⟨comp' p r c, a, convE e, s⟩ =|> ⟨c , conv v, convE e, s⟩. 150 | 151 | (** Setup the induction proof *) 152 | 153 | Proof. 154 | intros F E. 155 | generalize dependent c. 156 | generalize dependent a. 157 | generalize dependent s. 158 | generalize dependent r. 159 | induction E;intros. 160 | 161 | (** Calculation of the compiler *) 162 | 163 | (** - [Val n ⇓[e] Num n]: *) 164 | 165 | begin 166 | ⟨c, Num' n , convE e, s⟩. 167 | <== { apply vm_load } 168 | ⟨LOAD n c, a, convE e, s⟩. 169 | []. 170 | 171 | (** - [Add x y ⇓[e] Num (m + n)]: *) 172 | 173 | begin 174 | ⟨c, Num' (m + n), convE e, s⟩. 175 | ⊑ { auto } 176 | ⟨c, Num' (m + n), convE e, s[r:=VAL (Num' m)]⟩ . 177 | <== { apply vm_add } 178 | ⟨ADD r c, Num' n, convE e, s[r:=VAL (Num' m)]⟩ . 179 | <|= { apply IHE2 } 180 | ⟨comp' y (next r) (ADD r c), Num' m, convE e, s[r:=VAL (Num' m)]⟩ . 181 | <== { apply vm_store } 182 | ⟨STORE r (comp' y (next r) (ADD r c)), Num' m, convE e, s⟩. 183 | <|= { apply IHE1 } 184 | ⟨comp' x r (STORE r (comp' y (next r) (ADD r c))), a, convE e, s⟩. 185 | []. 186 | 187 | (** - [Var i ⇓[e] v] *) 188 | 189 | begin 190 | ⟨c, conv v, convE e , s⟩. 191 | <== {apply vm_lookup; unfold convE; rewrite nth_map; rewr_assumption} 192 | ⟨LOOKUP i c, a , convE e, s ⟩. 193 | []. 194 | 195 | (** - [Abs x ⇓[e] Clo x e] *) 196 | 197 | begin 198 | ⟨c, Clo' (fun r' => comp' x (next r') (RET r')) (convE e), convE e, s ⟩. 199 | <== { apply vm_fun } 200 | ⟨FUN (fun r' => comp' x (next r') (RET r')) c, a, convE e, s ⟩. 201 | []. 202 | 203 | (** - [App x y ⇓[e] x''] *) 204 | 205 | begin 206 | ⟨c, conv x'', convE e, s ⟩. 207 | ⊑ { auto } 208 | ⟨c, conv x'', convE e, s[r:=CLO c (convE e)] ⟩. 209 | <== { apply vm_ret } 210 | ⟨RET r, conv x'', convE (y' :: e'), s[r:=CLO c (convE e)]⟩. 211 | <|= { apply IHE3 } 212 | ⟨comp' x' (next r) (RET r), Num' 0, convE (y' :: e'), s[r:=CLO c (convE e)]⟩. 213 | = {auto} 214 | ⟨comp' x' (next r) (RET r), Num' 0, conv y' :: convE e', s[r:=CLO c (convE e)]⟩. 215 | <== {apply_eq vm_call} 216 | ⟨CALL r c, conv y', convE e, s[r:=VAL (Clo' (fun r' => comp' x' (next r') (RET r')) (convE e'))]⟩. 217 | <|= { apply IHE2 } 218 | ⟨comp' y (next r) (CALL r c), (Clo' (fun r' => comp' x' (next r') (RET r')) (convE e')), convE e, s[r:=VAL (Clo' (fun r' => comp' x' (next r') (RET r')) (convE e'))]⟩. 219 | <== { apply vm_store } 220 | ⟨STORE r (comp' y (next r) (CALL r c)), (Clo' (fun r' => comp' x' (next r') (RET r')) (convE e')), convE e, s⟩. 221 | = {auto} 222 | ⟨STORE r (comp' y (next r) (CALL r c)), conv (Clo x' e'), convE e,s ⟩. 223 | <|= { apply IHE1 } 224 | ⟨comp' x r (STORE r (comp' y (next r) (CALL r c))), a, convE e, s ⟩. 225 | []. 226 | Qed. 227 | 228 | (** * Soundness *) 229 | 230 | Lemma determ_vm : determ VM. 231 | intros C C1 C2 V. induction V; intro V'; inversion V'; subst; congruence. 232 | Qed. 233 | 234 | 235 | Definition terminates (p : Expr) : Prop := exists r, p ⇓[nil] r. 236 | 237 | Theorem sound p a s C : freeFrom first s -> terminates p -> ⟨comp p, a, nil, s⟩ =>>! C -> 238 | exists v s', C = ⟨HALT , conv v, nil, s'⟩ /\ p ⇓[nil] v. 239 | Proof. 240 | unfold terminates. intros F T M. destruct T as [v T]. 241 | pose (spec p v nil first HALT a s F T) as H'. 242 | unfold Reach in *. repeat autodestruct. 243 | pose (determ_trc determ_vm) as D. 244 | unfold determ in D. inversion H0. subst. exists v. eexists. split. eapply D. apply M. split. 245 | unfold comp. 246 | simpl in *. apply H. intro Contra. destruct Contra. 247 | inversion H1. assumption. 248 | Qed. 249 | 250 | End Lambda. 251 | -------------------------------------------------------------------------------- /coq/ExceptionPartial.v: -------------------------------------------------------------------------------- 1 | (** Calculation of the simple arithmetic language with exceptions 2 | using explicit empty handler. *) 3 | 4 | Require Import List. 5 | Require Import Tactics. 6 | Require Import Coq.Program.Equality. 7 | Module Exception (Import mem : Memory). 8 | 9 | 10 | 11 | (** * Syntax *) 12 | 13 | Inductive Expr : Set := 14 | | Val : nat -> Expr 15 | | Add : Expr -> Expr -> Expr 16 | | Throw : Expr 17 | | Catch : Expr -> Expr -> Expr. 18 | 19 | (** * Semantics *) 20 | 21 | Fixpoint eval (x: Expr) : option nat := 22 | match x with 23 | | Val n => Some n 24 | | Add x1 x2 => match eval x1 with 25 | | Some m => match eval x2 with 26 | | Some n => Some (m + n) 27 | | None => None 28 | end 29 | | None => None 30 | end 31 | | Throw => None 32 | | Catch x h => match eval x with 33 | | Some n => Some n 34 | | None => eval h 35 | end 36 | end. 37 | 38 | (** * Compiler *) 39 | 40 | Inductive Code : Set := 41 | | HALT : Code 42 | | LOAD : nat -> Code -> Code 43 | | STORE : Reg -> Code -> Code 44 | | ADD : Reg -> Code -> Code 45 | | THROW : Code 46 | | MARK : Reg -> Code -> Code -> Code 47 | | UNMARK : Code -> Code. 48 | 49 | 50 | Fixpoint comp (x : Expr) (r : Reg) (c : Code) : Code := 51 | match x with 52 | | Val n => LOAD n c 53 | | Add e1 e2 => comp e1 r (STORE r (comp e2 (next r) (ADD r c))) 54 | | Throw => THROW 55 | | Catch e1 e2 => MARK r (comp e2 r c) (comp e1 (next r) (UNMARK c)) 56 | end. 57 | 58 | Definition compile (x : Expr) : Code := comp x first HALT. 59 | 60 | Definition Han : Set := option (Code * Reg). 61 | 62 | Inductive RVal : Set := 63 | | VAL : nat -> RVal 64 | | HAN : Han -> RVal. 65 | 66 | (** * Virtual Machine *) 67 | 68 | Inductive Conf' : Type := 69 | | conf : Code -> nat -> Han -> Conf' 70 | | fail : Han -> Conf'. 71 | 72 | Definition Conf : Type := Conf' * Mem RVal. 73 | 74 | Notation "⟨ x , y , z , s ⟩" := (conf x y z, s). 75 | Notation "⟪ x , s ⟫" := (fail x, s). 76 | 77 | Reserved Notation "x ==> y" (at level 80, no associativity). 78 | Inductive VM : Conf -> Conf -> Prop := 79 | | vm_load n a c s h : ⟨LOAD n c, a , h, s⟩ ==> ⟨c , n, h, s⟩ 80 | | vm_store c s a r h : ⟨STORE r c, a, h, s⟩ ==> ⟨c , a, h, s[r:=VAL a]⟩ 81 | | vm_add c s a r n h : s[r]= VAL n -> ⟨ADD r c, a , h, s⟩ ==> ⟨c , n + a, h, s⟩ 82 | | vm_throw a s h : ⟨THROW, a, h, s⟩ ==> ⟪h, s⟫ 83 | | vm_mark h r s a c c' : ⟨MARK r c' c, a, h, s⟩ ==> ⟨c, a, Some (c', r), s[r:= HAN h]⟩ 84 | | vm_unmark h h' r s a c : s[r] = HAN h' -> 85 | ⟨UNMARK c, a, Some (h,r), s⟩ ==> ⟨c, a, h', s⟩ 86 | | vm_fail s c r h : s[r] = HAN h -> ⟪ Some (c,r), s⟫ ==> ⟨c, 0, h, s⟩ 87 | where "x ==> y" := (VM x y). 88 | 89 | Inductive cle : Conf -> Conf -> Prop := 90 | | cle_mem f s s' : s ⊑ s' -> cle (f, s) (f, s'). 91 | 92 | Hint Constructors cle : core. 93 | 94 | 95 | (** * Calculation *) 96 | 97 | (** Boilerplate to import calculation tactics *) 98 | 99 | Module VM <: Machine. 100 | Definition Conf := Conf. 101 | Definition Pre := cle. 102 | Definition Rel := VM. 103 | Lemma monotone : monotonicity cle VM. 104 | prove_monotonicity. Qed. 105 | Lemma preorder : is_preorder cle. 106 | prove_preorder. Qed. 107 | End VM. 108 | 109 | Module VMCalc := Calculation mem VM. 110 | Import VMCalc. 111 | 112 | 113 | (** Specification of the compiler *) 114 | 115 | Theorem spec : forall e r c a h s, 116 | freeFrom r s -> 117 | ⟨comp e r c, a, h, s⟩ 118 | =|> 119 | match eval e with 120 | | Some n => ⟨c , n, h, s⟩ 121 | | None => ⟪h, s⟫ 122 | end. 123 | 124 | 125 | (** Setup the induction proof *) 126 | 127 | Proof. 128 | induction e;intros. 129 | 130 | (** Calculation of the compiler *) 131 | 132 | (** - [x = Val n]: *) 133 | 134 | begin 135 | match eval (Val n) with 136 | | Some n' => ⟨c , n' ,h, s⟩ 137 | | None => ⟪h, s⟫ 138 | end. 139 | = {auto} 140 | ⟨ c, n, h, s⟩. 141 | <== {apply vm_load} 142 | ⟨ LOAD n c, a, h, s ⟩. 143 | []. 144 | 145 | (** - [x = Add x1 x2]: *) 146 | 147 | begin 148 | match eval (Add e1 e2) with 149 | | Some n => ⟨c , n ,h, s⟩ 150 | | None => ⟪h, s⟫ 151 | end. 152 | = { auto } 153 | match eval e1 with 154 | | Some n => match eval e2 with 155 | | Some n' => ⟨c , n + n' ,h, s⟩ 156 | | None => ⟪h, s⟫ 157 | end 158 | | None => ⟪h, s⟫ 159 | end. 160 | ⊑ { auto } 161 | match eval e1 with 162 | | Some n => match eval e2 with 163 | | Some n' => ⟨c , n + n' ,h, s[r:=VAL n]⟩ 164 | | None => ⟪h, s[r:=VAL n]⟫ 165 | end 166 | | None => ⟪h, s⟫ 167 | end. 168 | <== { apply vm_add } 169 | match eval e1 with 170 | | Some n => match eval e2 with 171 | | Some n' => ⟨ADD r c , n' ,h, s[r:=VAL n]⟩ 172 | | None => ⟪h, s[r:=VAL n]⟫ 173 | end 174 | | None => ⟪h, s⟫ 175 | end. 176 | <|= { apply IHe2 } 177 | match eval e1 with 178 | | Some n => ⟨comp e2 (next r) (ADD r c) , n ,h, s[r:=VAL n]⟩ 179 | | None => ⟪h, s⟫ 180 | end. 181 | <== { apply vm_store } 182 | match eval e1 with 183 | | Some n => ⟨STORE r (comp e2 (next r) (ADD r c)) , n ,h, s⟩ 184 | | None => ⟪h, s⟫ 185 | end. 186 | <|= { apply IHe1 } 187 | ⟨comp e1 r (STORE r (comp e2 (next r) (ADD r c))), a,h, s⟩. 188 | []. 189 | 190 | (** - [x = Throw]: *) 191 | 192 | begin 193 | match eval Throw with 194 | | Some n => ⟨c , n ,h, s⟩ 195 | | None => ⟪h, s⟫ 196 | end. 197 | = { auto } 198 | ⟪h, s⟫. 199 | <== { apply vm_throw } 200 | ⟨THROW, a, h, s⟩. 201 | []. 202 | 203 | (** - [x = Catch x1 x2]: *) 204 | 205 | begin 206 | match eval (Catch e1 e2) with 207 | | Some n => ⟨c , n ,h, s⟩ 208 | | None => ⟪h, s⟫ 209 | end. 210 | = { auto } 211 | match eval e1 with 212 | | Some n => ⟨c , n ,h, s⟩ 213 | | None => match eval e2 with 214 | | Some n' => ⟨c , n' ,h, s⟩ 215 | | None => ⟪h, s⟫ 216 | end 217 | end. 218 | <|= { apply IHe2 } 219 | match eval e1 with 220 | | Some n => ⟨c , n ,h, s⟩ 221 | | None => ⟨comp e2 r c, 0 ,h, s⟩ 222 | end. 223 | ⊑ { eauto } 224 | match eval e1 with 225 | | Some n => ⟨c , n ,h, s⟩ 226 | | None => ⟨comp e2 r c , 0 ,h, s[r:= HAN h]⟩ 227 | end. 228 | <== {apply vm_fail} 229 | match eval e1 with 230 | | Some n => ⟨c, n, h, s⟩ 231 | | None => ⟪Some (comp e2 r c,r), s[r:= HAN h]⟫ 232 | end. 233 | ⊑ {auto} 234 | match eval e1 with 235 | | Some n => ⟨c, n, h, s[r:= HAN h]⟩ 236 | | None => ⟪Some (comp e2 r c,r), s[r:= HAN h]⟫ 237 | end. 238 | <== {eapply vm_unmark} 239 | match eval e1 with 240 | | Some n => ⟨UNMARK c, n, Some (comp e2 r c,r), s[r:= HAN h]⟩ 241 | | None => ⟪Some (comp e2 r c,r), s[r:= HAN h]⟫ 242 | end. 243 | <|= {apply IHe1} 244 | ⟨comp e1 (next r) (UNMARK c) , a ,Some (comp e2 r c,r), s[r:= HAN h]⟩. 245 | <== {apply vm_mark} 246 | ⟨MARK r (comp e2 r c) (comp e1 (next r) (UNMARK c)) , a ,h, s⟩. 247 | []. 248 | Qed. 249 | 250 | Definition empty := (@empty RVal). 251 | 252 | (** Specification of the whole compiler *) 253 | 254 | Theorem spec' e a h : ⟨compile e, a, h, empty⟩ =|> match eval e with 255 | | Some n => ⟨HALT, n, h, empty⟩ 256 | | None => ⟪h, empty⟫ 257 | end. 258 | Proof. 259 | begin 260 | match eval e with 261 | | Some n => ⟨HALT, n, h, empty⟩ 262 | | None => ⟪h, empty⟫ 263 | end. 264 | <|= {apply spec; apply empty_mem_free} 265 | ⟨comp e first HALT, a,h, empty⟩. 266 | []. 267 | Qed. 268 | 269 | 270 | (** * Soundness *) 271 | 272 | (** Since the VM is defined as a small step operational semantics, we *) 273 | (* have to prove that the VM is deterministic and does not get stuck in *) 274 | (* order to derive soundness from the above theorem. *) 275 | 276 | 277 | Lemma determ_vm : determ VM. 278 | intros C C1 C2 V. induction V; intro V'; inversion V'; subst;eauto. 279 | rewrite H in *; dist. rewrite H in H6. dist. rewrite H in H4. dist. 280 | Qed. 281 | 282 | Theorem sound x a C : 283 | ⟨compile x, a, None, empty⟩ =>>! C -> 284 | ((exists m n, C = ⟨HALT, n, None, m⟩ /\ eval x = Some n) \/ (exists m, C = ⟪None , m⟫ /\ eval x = None)). 285 | Proof. 286 | intros E. 287 | pose (spec' x a None). unfold Reach in *. repeat autodestruct. 288 | pose (determ_trc determ_vm) as D. 289 | unfold determ in D. inversion H0. subst. 290 | remember (eval x) as X. 291 | destruct X. 292 | - left. inversion H0. subst. inversion H1. subst. eexists. eexists. 293 | split. eapply D. eapply E. split. apply H. intro Contra. destruct Contra. 294 | inversion H3. reflexivity. 295 | - right. inversion H0. subst. inversion H1. subst. eexists. 296 | split. eapply D. eapply E. split. apply H. intro Contra. destruct Contra. 297 | inversion H3. reflexivity. 298 | Qed. 299 | 300 | End Exception. 301 | -------------------------------------------------------------------------------- /coq/LambdaTruncate.v: -------------------------------------------------------------------------------- 1 | (** Calculation for the lambda calculus + arithmetic using a call 2 | stack which only holds a bounded number of registers (via 3 | truncate). *) 4 | 5 | Require Import List. 6 | Require Import ListIndex. 7 | Require Import Tactics. 8 | Require Import Coq.Program.Equality. 9 | Module Lambda (Import mem : TruncMem). 10 | 11 | (** * Syntax *) 12 | 13 | Inductive Expr : Set := 14 | | Val : nat -> Expr 15 | | Add : Expr -> Expr -> Expr 16 | | Var : nat -> Expr 17 | | Abs : Expr -> Expr 18 | | App : Expr -> Expr -> Expr. 19 | 20 | (** * Semantics *) 21 | 22 | (** The evaluator for this language is given as follows (as in the 23 | paper): 24 | << 25 | type Env = [Value] 26 | data Value = Num Int | Fun (Value -> Value) 27 | 28 | 29 | eval :: Expr -> Env -> Value 30 | eval (Val n) e = Num n 31 | eval (Add x y) e = case eval x e of 32 | Num n -> case eval y e of 33 | Num m -> Num (n+m) 34 | eval (Var i) e = e !! i 35 | eval (Abs x) e = Fun (\v -> eval x (v:e)) 36 | eval (App x y) e = case eval x e of 37 | Fun f -> f (eval y e) 38 | >> 39 | After defunctionalisation and translation into relational form we 40 | obtain the semantics below. *) 41 | 42 | Inductive Value : Set := 43 | | Num : nat -> Value 44 | | Clo : Expr -> list Value -> Value. 45 | 46 | Definition Env := list Value. 47 | 48 | Reserved Notation "x ⇓[ e ] y" (at level 80, no associativity). 49 | 50 | Inductive eval : Expr -> Env -> Value -> Prop := 51 | | eval_val e n : Val n ⇓[e] Num n 52 | | eval_add e x y m n : x ⇓[e] Num m -> y ⇓[e] Num n -> Add x y ⇓[e] Num (m + n) 53 | | eval_var e i v : nth e i = Some v -> Var i ⇓[e] v 54 | | eval_abs e x : Abs x ⇓[e] Clo x e 55 | | eval_app e e' x x' x'' y y' : x ⇓[e] Clo x' e' -> y ⇓[e] y' -> x' ⇓[y' :: e'] x'' -> App x y ⇓[e] x'' 56 | where "x ⇓[ e ] y" := (eval x e y). 57 | 58 | (** * Compiler *) 59 | 60 | Inductive Code : Set := 61 | | LOAD : nat -> Code -> Code 62 | | ADD : Reg -> Code -> Code 63 | | STORE : Reg -> Code -> Code 64 | | LOOKUP : nat -> Code -> Code 65 | | RET : Code 66 | | CALL : Reg -> Code -> Code 67 | | FUN : Code -> Code -> Code 68 | | HALT : Code. 69 | 70 | Fixpoint comp' (e : Expr) (r : Reg) (c : Code) : Code := 71 | match e with 72 | | Val n => LOAD n c 73 | | Add x y => comp' x r (STORE r (comp' y (next r) (ADD r c))) 74 | | Var i => LOOKUP i c 75 | | App x y => comp' x r (STORE r (comp' y (next r) (CALL r c))) 76 | | Abs x => FUN (comp' x (next first) RET) c 77 | end. 78 | 79 | Definition comp (e : Expr) : Code := comp' e first HALT. 80 | 81 | (** * Virtual Machine *) 82 | 83 | Inductive Value' : Set := 84 | | Num' : nat -> Value' 85 | | Clo' : Code -> list Value' -> Value'. 86 | 87 | Definition Env' := list Value'. 88 | 89 | Inductive RVal : Set := 90 | | VAL : Value' -> RVal 91 | | CLO : Code -> Env' -> RVal 92 | . 93 | 94 | Definition empty := (@empty RVal). 95 | 96 | 97 | Definition Lam : Type := list (Mem RVal). 98 | 99 | Inductive Conf : Type := 100 | | conf : Code -> Value' -> Env' -> Lam -> Mem RVal -> Conf. 101 | 102 | Notation "⟨ c , a , e , k , s ⟩" := (conf c a e k s). 103 | 104 | Reserved Notation "x ==> y" (at level 80, no associativity). 105 | Inductive VM : Conf -> Conf -> Prop := 106 | | vm_load n c s a e k : ⟨LOAD n c, a, e, k, s⟩ ==> ⟨c, Num' n, e, k, s⟩ 107 | | vm_add c m n r s e k : s[r] = VAL (Num' m) -> ⟨ADD r c, Num' n, e, k, s⟩ 108 | ==> ⟨c, Num'(m + n), e, k, s⟩ 109 | | vm_store c v r s e k : ⟨STORE r c, v, e, k, s⟩ 110 | ==> ⟨c, v, e, k, s[r:=VAL v]⟩ 111 | | vm_lookup e i c v a s k : nth e i = Some v -> ⟨LOOKUP i c, a, e, k, s⟩ ==> ⟨c, v, e, k, s⟩ 112 | | vm_ret a c e e' s s' k : s[first] = CLO c e -> ⟨RET, a, e', s' :: k, s⟩ ==> ⟨c, a, e, k, s'⟩ 113 | | vm_call c c' e e' v r s k : 114 | s[r]=VAL (Clo' c' e') -> 115 | ⟨CALL r c, v, e, k,s⟩ ==> ⟨c', Num' 0, v :: e', truncate r s :: k, empty[first:=CLO c e]⟩ 116 | | vm_fun a c c' s e k : ⟨FUN c' c, a, e, k, s⟩ ==> ⟨c, Clo' c' e, e, k, s⟩ 117 | where "x ==> y" := (VM x y). 118 | 119 | (** Conversion functions from semantics to VM *) 120 | 121 | Fixpoint conv (v : Value) : Value' := 122 | match v with 123 | | Num n => Num' n 124 | | Clo x e => Clo' (comp' x (next first) RET) (map conv e) 125 | end. 126 | 127 | Definition convE : Env -> Env' := map conv. 128 | 129 | Inductive stackle : Lam -> Lam -> Prop := 130 | | stackle_empty : stackle nil nil 131 | | stackle_cons s s' k k' : s ⊑ s' -> stackle k k' -> stackle (s :: k) (s' :: k'). 132 | 133 | Hint Constructors stackle : memory. 134 | 135 | Lemma stackle_refl k : stackle k k. 136 | Proof. 137 | induction k; constructor; auto with memory. 138 | Qed. 139 | 140 | Lemma stackle_trans k1 k2 k3 : stackle k1 k2 -> stackle k2 k3 -> stackle k1 k3. 141 | Proof. 142 | intros L1. generalize k3. induction L1; intros k3' L2. assumption. inversion L2. subst. constructor; 143 | eauto with memory. 144 | Qed. 145 | 146 | Hint Resolve stackle_refl stackle_trans : core. 147 | 148 | Inductive cle : Conf -> Conf -> Prop := 149 | | cle_mem c a e k k' s s' : stackle k k' -> s ⊑ s' -> cle ⟨ c , a , e , k, s ⟩ ⟨ c , a , e , k', s' ⟩. 150 | 151 | Hint Constructors cle : core. 152 | 153 | Lemma rel_eq {T} {R : T -> T -> Prop} x y y' : R x y' -> y = y' -> R x y. 154 | Proof. intros. subst. auto. 155 | Qed . 156 | Lemma rel_eq' {T} {R : T -> T -> Prop} x x' y : R x' y -> x = x' -> R x y. 157 | Proof. intros. subst. auto. 158 | Qed . 159 | 160 | Ltac apply_eq t := eapply rel_eq; [apply t | repeat rewrite set_set; auto]. 161 | 162 | 163 | (** * Calculation *) 164 | 165 | (** Boilerplate to import calculation tactics *) 166 | 167 | Module VM <: Machine. 168 | Definition Conf := Conf. 169 | Definition Pre := cle. 170 | Definition Rel := VM. 171 | Lemma monotone : monotonicity cle VM. 172 | prove_monotonicity1; 173 | try (match goal with [H : stackle (_ :: _) _ |- _] => inversion H end) 174 | ; prove_monotonicity2. 175 | Qed. 176 | Lemma preorder : is_preorder cle. 177 | prove_preorder. Qed. 178 | End VM. 179 | 180 | 181 | Module VMCalc := Calculation mem VM. 182 | Import VMCalc. 183 | 184 | 185 | 186 | (** Specification of the compiler *) 187 | 188 | Theorem spec p v e r c a s k : 189 | freeFrom r s -> p ⇓[e] v -> 190 | ⟨comp' p r c, a, convE e, k, s⟩ =|> ⟨c , conv v, convE e, k, s⟩. 191 | 192 | (** Setup the induction proof *) 193 | 194 | Proof. 195 | intros F E. 196 | generalize dependent c. 197 | generalize dependent a. 198 | generalize dependent s. 199 | generalize dependent r. 200 | generalize dependent k. 201 | induction E;intros. 202 | 203 | (** Calculation of the compiler *) 204 | 205 | (** - [Val n ⇓[e] Num n]: *) 206 | 207 | begin 208 | ⟨c, Num' n , convE e, k, s⟩. 209 | <== { apply vm_load } 210 | ⟨LOAD n c, a, convE e, k, s⟩. 211 | []. 212 | 213 | (** - [Add x y ⇓[e] Num (m + n)]: *) 214 | 215 | begin 216 | ⟨c, Num' (m + n), convE e, k, s⟩. 217 | ⊑ {auto} 218 | ⟨c, Num' (m + n), convE e, k, s[r:=VAL (Num' m)]⟩ . 219 | <== { apply vm_add } 220 | ⟨ADD r c, Num' n, convE e, k, s[r:=VAL (Num' m)]⟩ . 221 | <|= { apply IHE2 } 222 | ⟨comp' y (next r) (ADD r c), Num' m, convE e, k, s[r:=VAL (Num' m)]⟩ . 223 | <== { apply vm_store } 224 | ⟨STORE r (comp' y (next r) (ADD r c)), Num' m, convE e, k, s⟩. 225 | <|= { apply IHE1 } 226 | ⟨comp' x r (STORE r (comp' y (next r) (ADD r c))), a, convE e, k, s⟩. 227 | []. 228 | 229 | (** - [Var i ⇓[e] v] *) 230 | 231 | begin 232 | ⟨c, conv v, convE e , k, s⟩. 233 | <== {apply vm_lookup; unfold convE; rewrite nth_map; rewr_assumption} 234 | ⟨LOOKUP i c, a , convE e, k, s ⟩. 235 | []. 236 | 237 | (** - [Abs x ⇓[e] Clo x e] *) 238 | 239 | begin 240 | ⟨c, Clo' (comp' x (next first) RET) (convE e), convE e, k, s ⟩. 241 | <== { apply vm_fun } 242 | ⟨FUN (comp' x (next first) RET) c, a, convE e, k, s ⟩. 243 | []. 244 | 245 | (** - [App x y ⇓[e] x''] *) 246 | 247 | begin 248 | ⟨c, conv x'', convE e, k, s ⟩. 249 | <== { apply vm_ret } 250 | ⟨RET, conv x'', convE (y' :: e'), s :: k, empty[first:=CLO c (convE e)]⟩. 251 | <|= {apply IHE3} 252 | ⟨comp' x' (next first) RET, Num' 0, convE (y' :: e'), s :: k, empty[first:=CLO c (convE e)]⟩. 253 | = {auto} 254 | ⟨comp' x' (next first) RET, Num' 0, conv y' :: convE e', s::k, empty[first:=CLO c (convE e)]⟩. 255 | <== {apply_eq vm_call;try rewrite truncate_set} 256 | ⟨CALL r c, conv y', convE e, k, s[r:=VAL (Clo' (comp' x' (next first) RET) (convE e'))]⟩. 257 | <|= {apply IHE2} 258 | ⟨comp' y (next r) (CALL r c), (Clo' (comp' x' (next first) RET) (convE e')), convE e, k, s[r:=VAL (Clo' (comp' x' (next first) RET) (convE e'))]⟩. 259 | <== { apply vm_store } 260 | ⟨STORE r (comp' y (next r) (CALL r c)), (Clo' (comp' x' (next first) RET) (convE e')), convE e, k, s⟩. 261 | = {auto} 262 | ⟨STORE r (comp' y (next r) (CALL r c)), conv (Clo x' e'), convE e, k, s ⟩. 263 | <|= { apply IHE1 } 264 | ⟨comp' x r (STORE r (comp' y (next r) (CALL r c))), a, convE e,k, s ⟩. 265 | []. 266 | Qed. 267 | 268 | (** * Soundness *) 269 | 270 | Lemma determ_vm : determ VM. 271 | intros C C1 C2 V. induction V; intro V'; inversion V'; subst; congruence. 272 | Qed. 273 | 274 | 275 | Definition terminates (p : Expr) : Prop := exists r, p ⇓[nil] r. 276 | 277 | Theorem sound p a s C : freeFrom first s -> terminates p -> ⟨comp p, a, nil, nil, s⟩ =>>! C -> 278 | exists v s', C = ⟨HALT , conv v, nil, nil, s'⟩ /\ p ⇓[nil] v. 279 | Proof. 280 | unfold terminates. intros F T M. destruct T as [v T]. 281 | pose (spec p v nil first HALT a s nil F T) as H'. 282 | unfold Reach in *. repeat autodestruct. 283 | pose (determ_trc determ_vm) as D. 284 | unfold determ in D. inversion H0. inversion H7. subst. exists v. eexists. split. eapply D. apply M. split. 285 | unfold comp. 286 | simpl in *. apply H. intro Contra. destruct Contra. 287 | inversion H1. assumption. 288 | Qed. 289 | 290 | End Lambda. 291 | -------------------------------------------------------------------------------- /coq/Exception.v: -------------------------------------------------------------------------------- 1 | (** Calculation of the simple arithmetic language with exceptions 2 | (using CRASH instruction to signal uncaught exceptions). *) 3 | 4 | Require Import List. 5 | Require Import Tactics. 6 | Require Import Coq.Program.Equality. 7 | Module Exception (Import mem : Memory). 8 | 9 | 10 | 11 | (** * Syntax *) 12 | 13 | Inductive Expr : Set := 14 | | Val : nat -> Expr 15 | | Add : Expr -> Expr -> Expr 16 | | Throw : Expr 17 | | Catch : Expr -> Expr -> Expr. 18 | 19 | (** * Semantics *) 20 | 21 | Fixpoint eval (x: Expr) : option nat := 22 | match x with 23 | | Val n => Some n 24 | | Add x y => match eval x with 25 | | Some n => match eval y with 26 | | Some n' => Some (n + n') 27 | | None => None 28 | end 29 | | None => None 30 | end 31 | | Throw => None 32 | | Catch x y => match eval x with 33 | | Some n => Some n 34 | | None => eval y 35 | end 36 | end. 37 | 38 | (** * Compiler *) 39 | 40 | Inductive Code : Set := 41 | | HALT : Code 42 | | LOAD : nat -> Code -> Code 43 | | STORE : Reg -> Code -> Code 44 | | ADD : Reg -> Code -> Code 45 | | THROW : Code 46 | | MARK : Reg -> Code -> Code -> Code 47 | | UNMARK : Code -> Code 48 | | CRASH : Code. 49 | 50 | 51 | Fixpoint comp (x : Expr) (r : Reg) (c : Code) : Code := 52 | match x with 53 | | Val n => LOAD n c 54 | | Add x y => comp x r (STORE r (comp y (next r) (ADD r c))) 55 | | Throw => THROW 56 | | Catch x y => MARK r (comp y r c) (comp x (next r) (UNMARK c)) 57 | end. 58 | 59 | Definition compile (x : Expr) : Code := comp x first HALT. 60 | 61 | Definition Han : Set := Code * Reg. 62 | 63 | Inductive val : Set := 64 | | VAL : nat -> val 65 | | HAN : Han -> val. 66 | 67 | (** * Virtual Machine *) 68 | 69 | Inductive Conf' : Type := 70 | | conf : Code -> nat -> Han -> Conf' 71 | | fail : Han -> Conf'. 72 | 73 | Definition Conf : Type := Conf' * Mem val. 74 | 75 | Notation "⟨ x , y , z , m ⟩" := (conf x y z, m). 76 | Notation "⟪ x , m ⟫" := (fail x, m). 77 | 78 | Reserved Notation "x ==> y" (at level 80, no associativity). 79 | Inductive VM : Conf -> Conf -> Prop := 80 | | vm_load n a c m h : 81 | ⟨LOAD n c, a , h, m⟩ ==> ⟨c , n, h, m⟩ 82 | | vm_store c m a r h : 83 | ⟨STORE r c, a, h, m⟩ ==> ⟨c , a, h, m[r:=VAL a]⟩ 84 | | vm_add c m a r n h : m[r]= VAL n -> 85 | ⟨ADD r c, a , h, m⟩ ==> ⟨c , n + a, h, m⟩ 86 | | vm_throw a m h : ⟨THROW, a, h, m⟩ ==> ⟪h, m⟫ 87 | | vm_mark h r m a c c' : 88 | ⟨MARK r c' c, a, h, m⟩ ==> ⟨c, a, (c', r), m[r:= HAN h]⟩ 89 | | vm_unmark h h' r m a c : m[r] = HAN h' -> 90 | ⟨UNMARK c, a, (h,r), m⟩ ==> ⟨c, a, h', m⟩ 91 | | vm_fail m c r h : m[r] = HAN h -> 92 | ⟪(c,r), m⟫ ==> ⟨c, 0, h, m⟩ 93 | where "x ==> y" := (VM x y). 94 | 95 | Inductive cle : Conf -> Conf -> Prop := 96 | | cle_mem f m m' : m ⊑ m' -> cle (f, m) (f, m'). 97 | 98 | Hint Constructors cle : core. 99 | 100 | 101 | (** * Calculation *) 102 | 103 | (** Boilerplate to import calculation tactics *) 104 | 105 | Module VM <: Machine. 106 | Definition Conf := Conf. 107 | Definition Pre := cle. 108 | Definition Rel := VM. 109 | Lemma monotone : monotonicity cle VM. 110 | prove_monotonicity. Qed. 111 | Lemma preorder : is_preorder cle. 112 | prove_preorder. Qed. 113 | End VM. 114 | 115 | Module VMCalc := Calculation mem VM. 116 | Import VMCalc. 117 | 118 | 119 | (** Specification of the compiler *) 120 | 121 | Theorem spec : forall e r c a h m , 122 | freeFrom r m -> 123 | ⟨comp e r c, a, h, m⟩ 124 | =|> 125 | match eval e with 126 | | Some n => ⟨c , n, h, m⟩ 127 | | None => ⟪h, m⟫ 128 | end. 129 | 130 | 131 | (** Setup the induction proof *) 132 | 133 | Proof. 134 | induction e;intros. 135 | 136 | (** Calculation of the compiler *) 137 | 138 | (** - [x = Val n]: *) 139 | 140 | begin 141 | match eval (Val n) with 142 | | Some n' => ⟨c , n' ,h, m⟩ 143 | | None => ⟪h, m⟫ 144 | end. 145 | = {auto} 146 | ⟨ c, n, h, m⟩. 147 | <== {apply vm_load} 148 | ⟨ LOAD n c, a, h, m ⟩. 149 | []. 150 | 151 | (** - [x = Add e1 e2]: *) 152 | 153 | begin 154 | match eval (Add e1 e2) with 155 | | Some n => ⟨c , n ,h, m⟩ 156 | | None => ⟪h, m⟫ 157 | end. 158 | = { auto } 159 | match eval e1 with 160 | | Some n => match eval e2 with 161 | | Some n' => ⟨c , n + n' ,h, m⟩ 162 | | None => ⟪h, m⟫ 163 | end 164 | | None => ⟪h, m⟫ 165 | end. 166 | ⊑ { auto } 167 | match eval e1 with 168 | | Some n => match eval e2 with 169 | | Some n' => ⟨c , n + n' ,h, m[r:=VAL n]⟩ 170 | | None => ⟪h, m[r:=VAL n]⟫ 171 | end 172 | | None => ⟪h, m⟫ 173 | end. 174 | <== { apply vm_add } 175 | match eval e1 with 176 | | Some n => match eval e2 with 177 | | Some n' => ⟨ADD r c , n' ,h, m[r:=VAL n]⟩ 178 | | None => ⟪h, m[r:=VAL n]⟫ 179 | end 180 | | None => ⟪h, m⟫ 181 | end. 182 | <|= { apply IHe2 } 183 | match eval e1 with 184 | | Some n => ⟨comp e2 (next r) (ADD r c) , n ,h, m[r:=VAL n]⟩ 185 | | None => ⟪h, m⟫ 186 | end. 187 | <== { apply vm_store } 188 | match eval e1 with 189 | | Some n => ⟨STORE r (comp e2 (next r) (ADD r c)) , n ,h, m⟩ 190 | | None => ⟪h, m⟫ 191 | end. 192 | <|= { apply IHe1 } 193 | ⟨comp e1 r (STORE r (comp e2 (next r) (ADD r c))), a,h, m⟩. 194 | []. 195 | 196 | (** - [x = Throw]: *) 197 | 198 | begin 199 | match eval Throw with 200 | | Some n => ⟨c , n ,h, m⟩ 201 | | None => ⟪h, m⟫ 202 | end. 203 | = { auto } 204 | ⟪h, m⟫. 205 | <== { apply vm_throw } 206 | ⟨THROW, a, h, m⟩. 207 | []. 208 | 209 | (** - [x = Catch e1 e2]: *) 210 | 211 | begin 212 | match eval (Catch e1 e2) with 213 | | Some n => ⟨c , n ,h, m⟩ 214 | | None => ⟪h, m⟫ 215 | end. 216 | = { auto } 217 | match eval e1 with 218 | | Some n => ⟨c , n ,h, m⟩ 219 | | None => match eval e2 with 220 | | Some n' => ⟨c , n' ,h, m⟩ 221 | | None => ⟪h, m⟫ 222 | end 223 | end. 224 | <|= { apply IHe2 } 225 | match eval e1 with 226 | | Some n => ⟨c , n ,h, m⟩ 227 | | None => ⟨comp e2 r c, 0 ,h, m⟩ 228 | end. 229 | ⊑ { eauto } 230 | match eval e1 with 231 | | Some n => ⟨c , n ,h, m⟩ 232 | | None => ⟨comp e2 r c , 0 ,h, m[r:= HAN h]⟩ 233 | end. 234 | <== {apply vm_fail} 235 | match eval e1 with 236 | | Some n => ⟨c, n, h, m⟩ 237 | | None => ⟪(comp e2 r c,r), m[r:= HAN h]⟫ 238 | end. 239 | ⊑ {auto} 240 | match eval e1 with 241 | | Some n => ⟨c, n, h, m[r:= HAN h]⟩ 242 | | None => ⟪(comp e2 r c,r), m[r:= HAN h]⟫ 243 | end. 244 | <== {eapply vm_unmark} 245 | match eval e1 with 246 | | Some n => ⟨UNMARK c, n, (comp e2 r c,r), m[r:= HAN h]⟩ 247 | | None => ⟪(comp e2 r c,r), m[r:= HAN h]⟫ 248 | end. 249 | <|= {apply IHe1} 250 | ⟨comp e1 (next r) (UNMARK c) , a ,(comp e2 r c,r), m[r:= HAN h]⟩. 251 | <== {apply vm_mark} 252 | ⟨MARK r (comp e2 r c) (comp e1 (next r) (UNMARK c)) , a ,h, m⟩. 253 | []. 254 | Qed. 255 | 256 | Definition empty := (@empty val). 257 | 258 | (** Specification of the whole compiler *) 259 | 260 | Theorem spec' e a h : ⟨compile e, a, h, empty⟩ =|> match eval e with 261 | | Some n => ⟨HALT, n, h, empty⟩ 262 | | None => ⟪h, empty⟫ 263 | end. 264 | Proof. 265 | begin 266 | match eval e with 267 | | Some n => ⟨HALT, n, h, empty⟩ 268 | | None => ⟪h, empty⟫ 269 | end. 270 | <|= {apply spec; apply empty_mem_free} 271 | ⟨comp e first HALT, a,h, empty⟩. 272 | []. 273 | Qed. 274 | 275 | 276 | (** * Soundness *) 277 | 278 | (** Since the VM is defined as a small step operational semantics, we *) 279 | (* have to prove that the VM is deterministic and does not get stuck in *) 280 | (* order to derive soundness from the above theorem. *) 281 | 282 | 283 | Lemma determ_vm : determ VM. 284 | intros C C1 C2 V. induction V; intro V'; inversion V'; subst;eauto. 285 | rewrite H in *; dist. rewrite H in H6. dist. rewrite H in H4. dist. 286 | Qed. 287 | 288 | Theorem sound' x a r C : 289 | ⟨compile x, a, (CRASH, r), empty⟩ =>>! C -> 290 | ((exists m n, C = ⟨HALT, n, (CRASH, r), m⟩ /\ eval x = Some n) 291 | \/ ((forall n h m, C <> ⟨HALT, n, h, m⟩) /\ eval x = None)). 292 | Proof. 293 | intros E. 294 | pose (spec' x a (CRASH, r)). unfold Reach in *. repeat autodestruct. 295 | pose (determ_trc determ_vm) as D. 296 | unfold determ in D. inversion H0. subst. 297 | remember (eval x) as X. 298 | destruct X. 299 | - left. inversion H0. subst. inversion H1. subst. eexists. eexists. 300 | split. eapply D. eapply E. split. apply H. intro Contra. destruct Contra. 301 | inversion H3. reflexivity. 302 | - right. inversion H0. subst. inversion H1. subst. 303 | remember (get r m') as G. split;auto. 304 | destruct G. destruct v. 305 | + assert (C = ⟪ (CRASH, r), m' ⟫) as R. eapply D. eapply E. split. auto. intro Contra. destruct Contra. 306 | inversion H3. rewrite H9 in HeqG. inversion HeqG. subst. intros. intro Contra. inversion Contra. 307 | + assert (C = ⟨ CRASH, 0, h, m' ⟩) as R. 308 | eapply D. eapply E. split. eapply trc_step_trans. apply H. econstructor. eauto. intro Contra. destruct Contra. 309 | inversion H3. subst. intros. intro Contra. inversion Contra. 310 | + assert (C = ⟪ (CRASH, r), m' ⟫) as R. eapply D. eapply E. split. auto. intro Contra. destruct Contra. 311 | inversion H3. rewrite H9 in HeqG. inversion HeqG. subst. intros. intro Contra. inversion Contra. 312 | Qed. 313 | 314 | 315 | End Exception. 316 | -------------------------------------------------------------------------------- /coq/Lambda.v: -------------------------------------------------------------------------------- 1 | (** Calculation for the lambda calculus + arithmetic using a call 2 | stack for saving registers before a call. *) 3 | 4 | Require Import List. 5 | Require Import ListIndex. 6 | Require Import Tactics. 7 | Require Import Coq.Program.Equality. 8 | Module Lambda (Import mem : Memory). 9 | 10 | 11 | (** * Syntax *) 12 | 13 | Inductive Expr : Set := 14 | | Val : nat -> Expr 15 | | Add : Expr -> Expr -> Expr 16 | | Var : nat -> Expr 17 | | Abs : Expr -> Expr 18 | | App : Expr -> Expr -> Expr. 19 | 20 | (** * Semantics *) 21 | 22 | (** The evaluator for this language is given as follows (as in the 23 | paper): 24 | << 25 | type Env = [Value] 26 | data Value = Num Int | Fun (Value -> Value) 27 | 28 | 29 | eval :: Expr -> Env -> Value 30 | eval (Val n) e = Num n 31 | eval (Add x y) e = case eval x e of 32 | Num n -> case eval y e of 33 | Num m -> Num (n+m) 34 | eval (Var i) e = e !! i 35 | eval (Abs x) e = Fun (\v -> eval x (v:e)) 36 | eval (App x y) e = case eval x e of 37 | Fun f -> f (eval y e) 38 | >> 39 | After defunctionalisation and translation into relational form we 40 | obtain the semantics below. *) 41 | 42 | Inductive Value : Set := 43 | | Num : nat -> Value 44 | | Clo : Expr -> list Value -> Value. 45 | 46 | Definition Env := list Value. 47 | 48 | Reserved Notation "x ⇓[ e ] y" (at level 80, no associativity). 49 | 50 | Inductive eval : Expr -> Env -> Value -> Prop := 51 | | eval_val e n : Val n ⇓[e] Num n 52 | | eval_add e x y n n' : x ⇓[e] Num n -> y ⇓[e] Num n' -> Add x y ⇓[e] Num (n + n') 53 | | eval_var e i v : nth e i = Some v -> Var i ⇓[e] v 54 | | eval_abs e x : Abs x ⇓[e] Clo x e 55 | | eval_app e e' x x' x'' y y' : x ⇓[e] Clo x' e' -> y ⇓[e] y' -> x' ⇓[y' :: e'] x'' -> App x y ⇓[e] x'' 56 | where "x ⇓[ e ] y" := (eval x e y). 57 | 58 | (** * Compiler *) 59 | 60 | Inductive Code : Set := 61 | | HALT : Code 62 | | LOAD : nat -> Code -> Code 63 | | LOOKUP : nat -> Code -> Code 64 | | STORE : Reg -> Code -> Code 65 | | ADD : Reg -> Code -> Code 66 | | STC : Reg -> Code -> Code 67 | | APP : Reg -> Code -> Code 68 | | ABS : Code -> Code -> Code 69 | | RET : Code. 70 | 71 | Fixpoint comp (e : Expr) (r : Reg) (c : Code) : Code := 72 | match e with 73 | | Val n => LOAD n c 74 | | Var i => LOOKUP i c 75 | | Add x y => comp x r (STORE r (comp y (next r) (ADD r c))) 76 | | App x y => comp x r (STC r (comp y (next r) (APP r c))) 77 | | Abs x => ABS (comp x (next first) RET) c 78 | end. 79 | 80 | Definition compile (e : Expr) : Code := comp e first HALT. 81 | 82 | (** * Virtual Machine *) 83 | 84 | Inductive Value' : Set := 85 | | Num' : nat -> Value' 86 | | Clo' : Code -> list Value' -> Value'. 87 | 88 | Definition Env' := list Value'. 89 | 90 | Inductive val : Set := 91 | | NUM : nat -> val 92 | | CLO : Code -> Env' -> val. 93 | 94 | Definition empty := (@empty val). 95 | 96 | 97 | Definition Lam : Type := list (Mem val). 98 | 99 | Inductive Conf : Type := 100 | | conf : Code -> Value' -> Env' -> Lam -> Mem val -> Conf. 101 | 102 | Notation "⟨ c , a , e , s , m ⟩" := (conf c a e s m). 103 | 104 | Reserved Notation "x ==> y" (at level 80, no associativity). 105 | Inductive VM : Conf -> Conf -> Prop := 106 | | vm_load n c m a e s : 107 | ⟨LOAD n c, a, e, s, m⟩ ==> ⟨c, Num' n, e, s, m⟩ 108 | | vm_add c a n r m e s : m[r] = NUM n -> 109 | ⟨ADD r c, Num' a, e, s, m⟩ ==> ⟨c, Num'(n + a), e, s, m⟩ 110 | | vm_store c n r m e s : 111 | ⟨STORE r c, Num' n, e, s, m⟩ ==> ⟨c, Num' n, e, s, m[r:=NUM n]⟩ 112 | | vm_stc c c' e' r m e s : 113 | ⟨STC r c, Clo' c' e', e, s, m⟩ ==> ⟨c, Clo' c' e', e, s, m[r:=CLO c' e']⟩ 114 | | vm_lookup e i c v a m s : nth e i = Some v -> 115 | ⟨LOOKUP i c, a, e, s, m⟩ ==> ⟨c, v, e, s, m⟩ 116 | | vm_ret a c e e' m m' s : m[first] = CLO c e -> 117 | ⟨RET, a, e', m' :: s, m⟩ ==> ⟨c, a, e, s, m'⟩ 118 | | vm_call c c' e e' v r m s : m[r]=CLO c' e' -> 119 | ⟨APP r c, v, e, s, m⟩ ==> ⟨c', v, v :: e', m :: s, empty[first:=CLO c e]⟩ 120 | | vm_fun a c c' m e s : 121 | ⟨ABS c' c, a, e, s, m⟩ ==> ⟨c, Clo' c' e, e, s, m⟩ 122 | where "x ==> y" := (VM x y). 123 | 124 | (** Conversion functions from semantics to VM *) 125 | 126 | Fixpoint conv (v : Value) : Value' := 127 | match v with 128 | | Num n => Num' n 129 | | Clo x e => Clo' (comp x (next first) RET) (map conv e) 130 | end. 131 | 132 | Definition convE : Env -> Env' := map conv. 133 | 134 | Inductive stackle : Lam -> Lam -> Prop := 135 | | stackle_empty : stackle nil nil 136 | | stackle_cons m m' s s' : m ⊑ m' -> stackle s s' -> stackle (m :: s) (m' :: s'). 137 | 138 | Hint Constructors stackle : memory. 139 | 140 | Lemma stackle_refl s : stackle s s. 141 | Proof. 142 | induction s; constructor; auto with memory. 143 | Qed. 144 | 145 | Lemma stackle_trans s1 s2 s3 : stackle s1 s2 -> stackle s2 s3 -> stackle s1 s3. 146 | Proof. 147 | intros L1. generalize s3. induction L1; intros s3' L2. assumption. inversion L2. subst. constructor; 148 | eauto with memory. 149 | Qed. 150 | 151 | Hint Resolve stackle_refl stackle_trans : core. 152 | 153 | Inductive cle : Conf -> Conf -> Prop := 154 | | cle_mem c a e s s' m m' : stackle s s' -> m ⊑ m' -> cle ⟨ c , a , e , s, m ⟩ ⟨ c , a , e , s', m' ⟩. 155 | 156 | Hint Constructors cle : core. 157 | 158 | Lemma rel_eq {T} {R : T -> T -> Prop} x y y' : R x y' -> y = y' -> R x y. 159 | Proof. intros. subst. auto. 160 | Qed . 161 | Lemma rel_eq' {T} {R : T -> T -> Prop} x x' y : R x' y -> x = x' -> R x y. 162 | Proof. intros. subst. auto. 163 | Qed . 164 | 165 | 166 | (** * Calculation *) 167 | 168 | (** Boilerplate to import calculation tactics *) 169 | 170 | Module VM <: Machine. 171 | Definition Conf := Conf. 172 | Definition Pre := cle. 173 | Definition Rel := VM. 174 | Lemma monotone : monotonicity cle VM. 175 | prove_monotonicity1; 176 | try (match goal with [H : stackle (_ :: _) _ |- _] => inversion H end) 177 | ; prove_monotonicity2. 178 | Qed. 179 | Lemma preorder : is_preorder cle. 180 | prove_preorder. Qed. 181 | End VM. 182 | 183 | 184 | Module VMCalc := Calculation mem VM. 185 | Import VMCalc. 186 | 187 | 188 | 189 | (** Specification of the compiler *) 190 | 191 | Theorem spec p v e r c a m s : 192 | freeFrom r m -> p ⇓[e] v -> 193 | ⟨comp p r c, a, convE e, s, m⟩ =|> ⟨c , conv v, convE e, s, m⟩. 194 | 195 | (** Setup the induction proof *) 196 | 197 | Proof. 198 | intros F E. 199 | generalize dependent c. 200 | generalize dependent a. 201 | generalize dependent m. 202 | generalize dependent r. 203 | generalize dependent s. 204 | induction E;intros. 205 | 206 | (** Calculation of the compiler *) 207 | 208 | (** - [Val n ⇓[e] Num n]: *) 209 | 210 | begin 211 | ⟨c, Num' n , convE e, s, m⟩. 212 | <== { apply vm_load } 213 | ⟨LOAD n c, a, convE e, s, m⟩. 214 | []. 215 | 216 | (** - [Add x y ⇓[e] Num (n + n')]: *) 217 | 218 | begin 219 | ⟨c, Num' (n + n'), convE e, s, m⟩. 220 | ⊑ {auto} 221 | ⟨c, Num' (n + n'), convE e, s, m[r:=NUM n]⟩ . 222 | <== { apply vm_add } 223 | ⟨ADD r c, Num' n', convE e, s, m[r:=NUM n]⟩ . 224 | <|= { apply IHE2 } 225 | ⟨comp y (next r) (ADD r c), Num' n, convE e, s, m[r:=NUM n]⟩ . 226 | <== { apply vm_store } 227 | ⟨STORE r (comp y (next r) (ADD r c)), Num' n, convE e, s, m⟩. 228 | <|= { apply IHE1 } 229 | ⟨comp x r (STORE r (comp y (next r) (ADD r c))), a, convE e, s, m⟩. 230 | []. 231 | 232 | (** - [Var i ⇓[e] v] *) 233 | 234 | begin 235 | ⟨c, conv v, convE e , s, m⟩. 236 | <== {apply vm_lookup; unfold convE; rewrite nth_map; rewr_assumption} 237 | ⟨LOOKUP i c, a , convE e, s, m ⟩. 238 | []. 239 | 240 | (** - [Abs x ⇓[e] Clo x e] *) 241 | 242 | begin 243 | ⟨c, Clo' (comp x (next first) RET) (convE e), convE e, s, m ⟩. 244 | <== { apply vm_fun } 245 | ⟨ABS (comp x (next first) RET) c, a, convE e, s, m ⟩. 246 | []. 247 | 248 | (** - [App x y ⇓[e] x''] *) 249 | 250 | begin 251 | ⟨c, conv x'', convE e, s, m ⟩. 252 | <== { apply vm_ret } 253 | ⟨RET, conv x'', convE (y' :: e'), m :: s, empty[first:=CLO c (convE e)]⟩. 254 | <|= {apply IHE3} 255 | ⟨comp x' (next first) RET, conv y', convE (y' :: e'), m :: s, empty[first:=CLO c (convE e)]⟩. 256 | = {auto} 257 | ⟨comp x' (next first) RET, conv y', conv y' :: convE e', m::s, empty[first:=CLO c (convE e)]⟩. 258 | ⊑ {auto with memory} 259 | ⟨comp x' (next first) RET, conv y', conv y' :: convE e', m[r:=CLO (comp x' (next first) RET) (convE e')]::s, empty[first:=CLO c (convE e)]⟩. 260 | <== {apply vm_call} 261 | ⟨APP r c, conv y', convE e, s, m[r:=CLO (comp x' (next first) RET) (convE e')]⟩. 262 | <|= {apply IHE2} 263 | ⟨comp y (next r) (APP r c), (Clo' (comp x' (next first) RET) (convE e')), convE e, s, m[r:=CLO (comp x' (next first) RET) (convE e')]⟩. 264 | <== { apply vm_stc } 265 | ⟨STC r (comp y (next r) (APP r c)), (Clo' (comp x' (next first) RET) (convE e')), convE e, s, m⟩. 266 | = {auto} 267 | ⟨STC r (comp y (next r) (APP r c)), conv (Clo x' e'), convE e, s, m ⟩. 268 | <|= { apply IHE1 } 269 | ⟨comp x r (STC r (comp y (next r) (APP r c))), a, convE e, s, m ⟩. 270 | []. 271 | Qed. 272 | 273 | 274 | (** Specification of the whole compiler *) 275 | 276 | Theorem spec' x a (e : Env) s v : x ⇓[e] v -> 277 | ⟨compile x, a, convE e, s, empty⟩ =|> ⟨HALT , conv v , convE e, s, empty⟩. 278 | Proof. 279 | intros E. 280 | begin 281 | ⟨HALT , conv v , convE e, s, empty⟩. 282 | <|= {apply spec; auto using empty_mem_free} 283 | ⟨comp x first HALT, a, convE e, s, empty⟩. 284 | []. 285 | Qed. 286 | 287 | 288 | (** * Soundness *) 289 | 290 | Lemma determ_vm : determ VM. 291 | intros C C1 C2 V. induction V; intro V'; inversion V'; subst; congruence. 292 | Qed. 293 | 294 | 295 | Definition terminates (p : Expr) : Prop := exists r, p ⇓[nil] r. 296 | 297 | Theorem sound p a C : terminates p -> ⟨compile p, a, nil, nil, empty⟩ =>>! C -> 298 | exists v m, C = ⟨HALT , conv v, nil, nil, m⟩ /\ p ⇓[nil] v. 299 | Proof. 300 | unfold terminates. intros T M. destruct T as [v T]. 301 | pose (spec' p a nil nil v T) as H'. 302 | unfold Reach in *. repeat autodestruct. 303 | pose (determ_trc determ_vm) as D. 304 | unfold determ in D. inversion H0. inversion H7. subst. exists v. eexists. split. eapply D. apply M. split. 305 | unfold comp. 306 | simpl in *. apply H. intro Contra. destruct Contra. 307 | inversion H1. assumption. 308 | Qed. 309 | 310 | End Lambda. 311 | -------------------------------------------------------------------------------- /coq/LocalState.v: -------------------------------------------------------------------------------- 1 | (** Calculation of the simple arithmetic language with exceptions and 2 | local state.. *) 3 | 4 | Require Import List. 5 | Require Import Tactics. 6 | Require Import Coq.Program.Equality. 7 | Module LocalState (Import mem : Memory). 8 | 9 | 10 | (** * Syntax *) 11 | 12 | Inductive Expr : Set := 13 | | Val : nat -> Expr 14 | | Add : Expr -> Expr -> Expr 15 | | Throw : Expr 16 | | Catch : Expr -> Expr -> Expr 17 | | Get : Expr 18 | | Put : Expr -> Expr -> Expr. 19 | 20 | Definition state := nat. 21 | 22 | (** * Semantics *) 23 | 24 | Fixpoint eval (x: Expr) (q : state) : option (nat * state) := 25 | match x with 26 | | Val n => Some (n, q) 27 | | Add x1 x2 => match eval x1 q with 28 | | Some (m, q') => match eval x2 q' with 29 | | Some (n, q'') => Some (m + n, q'') 30 | | None => None 31 | end 32 | | None => None 33 | end 34 | | Throw => None 35 | | Catch x h => match eval x q with 36 | | Some (n, q') => Some (n, q') 37 | | None => eval h q 38 | end 39 | | Get => Some (q, q) 40 | | Put x y => match eval x q with 41 | | Some (n, q') => eval y n 42 | | None => None 43 | end 44 | end. 45 | 46 | (** * Compiler *) 47 | 48 | Inductive Code : Set := 49 | | LOAD : nat -> Code -> Code 50 | | ADD : Reg -> Code -> Code 51 | | STORE : Reg -> Code -> Code 52 | | THROW : Code 53 | | UNMARK : Code -> Code 54 | | MARK : Reg -> Code -> Code -> Code 55 | | READ : Code -> Code 56 | | WRITE : Code -> Code 57 | | HALT : Code. 58 | 59 | Fixpoint comp' (x : Expr) (r : Reg) (c : Code) : Code := 60 | match x with 61 | | Val n => LOAD n c 62 | | Add e1 e2 => comp' e1 r (STORE r (comp' e2 (next r) (ADD r c))) 63 | | Throw => THROW 64 | | Catch e1 e2 => MARK r (comp' e2 r c) (comp' e1 (next r) (UNMARK c)) 65 | | Get => READ c 66 | | Put x y => comp' x r (WRITE (comp' y r c)) 67 | end. 68 | 69 | Definition comp (x : Expr) : Code := comp' x first HALT. 70 | 71 | Definition Han := option Reg. 72 | 73 | Inductive RVal : Set := 74 | | VAL : nat -> RVal 75 | | HAN : Code -> Han -> state -> RVal. 76 | 77 | (** * Virtual Machine *) 78 | 79 | Inductive Conf' : Type := 80 | | conf : Code -> nat -> Han -> state -> Conf' 81 | | fail : Han -> Conf'. 82 | 83 | Definition Conf : Type := Conf' * Mem RVal. 84 | 85 | Notation "⟨ x , y , z , q , s ⟩" := (conf x y z q, s). 86 | Notation "⟪ x , s ⟫" := (fail x, s). 87 | 88 | Reserved Notation "x ==> y" (at level 80, no associativity). 89 | Inductive VM : Conf -> Conf -> Prop := 90 | | vm_load n a c s q p : ⟨LOAD n c, a , p, q, s⟩ ==> ⟨c , n, p, q, s⟩ 91 | | vm_add c s a r n q p : s[r]= VAL n -> ⟨ADD r c, a , p, q, s⟩ ==> ⟨c , n + a, p, q, s⟩ 92 | | vm_store c s a r q p : ⟨STORE r c, a, p, q, s⟩ ==> ⟨c , a, p, q, s[r:=VAL a]⟩ 93 | | vm_throw a s q p : ⟨THROW, a, p, q, s⟩ ==> ⟪p, s⟫ 94 | | vm_fail p p' s q c : s[p] = HAN c p' q -> ⟪ Some p, s⟫ ==> ⟨c, 0, p', q, s⟩ 95 | | vm_unmark p p' s a q q' c c' : s[p] = HAN c' p' q' -> 96 | ⟨UNMARK c, a, Some p, q, s⟩ ==> ⟨c, a, p', q, s⟩ 97 | | vm_mark p r s a c c' q : 98 | ⟨MARK r c' c, a, p, q, s⟩ ==> ⟨c, a, Some r, q, s[r:= HAN c' p q]⟩ 99 | | vm_read a c s q p : ⟨READ c, a , p, q, s⟩ ==> ⟨c , q, p, q, s⟩ 100 | | vm_write a c s q p : ⟨WRITE c, a , p, q, s⟩ ==> ⟨c , a, p, a, s⟩ 101 | where "x ==> y" := (VM x y). 102 | 103 | 104 | 105 | Inductive cle : Conf -> Conf -> Prop := 106 | | cle_mem f s s' : s ⊑ s' -> cle (f, s) (f, s'). 107 | 108 | Hint Constructors cle : core. 109 | 110 | 111 | (** * Calculation *) 112 | 113 | (** Boilerplate to import calculation tactics *) 114 | Module VM <: Machine. 115 | Definition Conf := Conf. 116 | Definition Pre := cle. 117 | Definition Rel := VM. 118 | Lemma monotone : monotonicity cle VM. 119 | prove_monotonicity. Qed. 120 | Lemma preorder : is_preorder cle. 121 | prove_preorder. Qed. 122 | End VM. 123 | 124 | Module VMCalc := Calculation mem VM. 125 | Import VMCalc. 126 | 127 | 128 | (** Specification of the compiler *) 129 | 130 | Theorem spec : forall e r c a p s q, 131 | freeFrom r s -> 132 | ⟨comp' e r c, a, p, q, s⟩ 133 | =|> 134 | match eval e q with 135 | | Some (n, q') => ⟨c , n, p, q', s⟩ 136 | | None => ⟪p, s⟫ 137 | end. 138 | 139 | 140 | (** Setup the induction proof *) 141 | 142 | Proof. 143 | induction e;intros. 144 | 145 | (** Calculation of the compiler *) 146 | 147 | (** - [x = Val n]: *) 148 | 149 | begin 150 | match eval (Val n) q with 151 | | Some (n', q') => ⟨c , n' ,p, q', s⟩ 152 | | None => ⟪p, s⟫ 153 | end. 154 | = {auto} 155 | ⟨ c, n, p, q, s⟩. 156 | <== {apply vm_load} 157 | ⟨ LOAD n c, a, p, q, s ⟩. 158 | []. 159 | 160 | (** - [x = Add x1 x2]: *) 161 | 162 | begin 163 | match eval (Add e1 e2) q with 164 | | Some (n, q') => ⟨c , n ,p, q', s⟩ 165 | | None => ⟪p, s⟫ 166 | end. 167 | = {auto} 168 | match eval e1 q with 169 | | Some (n, q') => match eval e2 q' with 170 | | Some (n', q'') => ⟨c , n + n' , p, q'', s⟩ 171 | | None => ⟪p, s⟫ 172 | end 173 | | None => ⟪p, s⟫ 174 | end. 175 | ⊑ { auto } 176 | match eval e1 q with 177 | | Some (n, q') => match eval e2 q' with 178 | | Some (n', q'') => ⟨c , n + n' ,p, q'', s[r:=VAL n]⟩ 179 | | None => ⟪p, s[r:=VAL n]⟫ 180 | end 181 | | None => ⟪p, s⟫ 182 | end. 183 | <== { apply vm_add } 184 | match eval e1 q with 185 | | Some (n, q') => match eval e2 q' with 186 | | Some (n', q'') => ⟨ADD r c , n' ,p, q'', s[r:=VAL n]⟩ 187 | | None => ⟪p, s[r:=VAL n]⟫ 188 | end 189 | | None => ⟪p, s⟫ 190 | end. 191 | <|= {apply IHe2} 192 | match eval e1 q with 193 | | Some (n, q') => ⟨comp' e2 (next r) (ADD r c) , n ,p, q', s[r:=VAL n]⟩ 194 | | None => ⟪p, s⟫ 195 | end. 196 | <== { apply vm_store } 197 | match eval e1 q with 198 | | Some (n, q') => ⟨STORE r (comp' e2 (next r) (ADD r c)) , n ,p, q', s⟩ 199 | | None => ⟪p, s⟫ 200 | end. 201 | <|= { apply IHe1 } 202 | ⟨comp' e1 r (STORE r (comp' e2 (next r) (ADD r c))), a,p, q, s⟩. 203 | []. 204 | 205 | (** - [x = Throw]: *) 206 | 207 | begin 208 | match eval Throw q with 209 | | Some (n, q') => ⟨c , n ,p, q', s⟩ 210 | | None => ⟪p, s⟫ 211 | end. 212 | = { auto } 213 | ⟪p, s⟫. 214 | <== { apply vm_throw } 215 | ⟨THROW, a, p, q, s⟩. 216 | []. 217 | 218 | (** - [x = Catch x1 x2]: *) 219 | 220 | begin 221 | match eval (Catch e1 e2) q with 222 | | Some (n, q') => ⟨c , n ,p, q', s⟩ 223 | | None => ⟪p, s⟫ 224 | end. 225 | = { auto } 226 | match eval e1 q with 227 | | Some (n, q') => ⟨c , n ,p, q', s⟩ 228 | | None => match eval e2 q with 229 | | Some (n', q'') => ⟨c , n' ,p, q'', s⟩ 230 | | None => ⟪p,s⟫ 231 | end 232 | end. 233 | <|= { apply IHe2 } 234 | match eval e1 q with 235 | | Some (n, q') => ⟨c, n ,p, q', s⟩ 236 | | None => ⟨comp' e2 r c , 0 ,p, q, s⟩ 237 | end. 238 | ⊑ { eauto } 239 | match eval e1 q with 240 | | Some (n, q') => ⟨c , n ,p, q', s⟩ 241 | | None => ⟨comp' e2 r c , 0 ,p, q, s[r:= HAN (comp' e2 r c) p q]⟩ 242 | end. 243 | <== {apply vm_fail} 244 | match eval e1 q with 245 | | Some (n, q') => ⟨c , n ,p, q', s⟩ 246 | | None => ⟪Some r, s[r:= HAN (comp' e2 r c) p q]⟫ 247 | end. 248 | ⊑ {auto} 249 | match eval e1 q with 250 | | Some (n, q') => ⟨c , n ,p, q', s[r:= HAN (comp' e2 r c) p q]⟩ 251 | | None => ⟪Some r, s[r:= HAN (comp' e2 r c) p q]⟫ 252 | end. 253 | <== {eapply vm_unmark} 254 | match eval e1 q with 255 | | Some (n, q') => ⟨UNMARK c , n,Some r, q', s[r:= HAN (comp' e2 r c) p q]⟩ 256 | | None => ⟪Some r, s[r:= HAN (comp' e2 r c) p q]⟫ 257 | end. 258 | <|= {apply IHe1} 259 | ⟨comp' e1 (next r) (UNMARK c) , a ,Some r, q, s[r:= HAN (comp' e2 r c) p q]⟩. 260 | <== {apply vm_mark} 261 | ⟨MARK r (comp' e2 r c) (comp' e1 (next r) (UNMARK c)) , a ,p, q, s⟩. 262 | []. 263 | 264 | 265 | begin 266 | match eval Get q with 267 | | Some (n, q') => ⟨c , n ,p, q', s⟩ 268 | | None => ⟪p, s⟫ 269 | end. 270 | = {auto} 271 | ⟨c , q ,p, q, s⟩. 272 | <== {apply vm_read} 273 | ⟨READ c , a ,p, q, s⟩. 274 | []. 275 | 276 | begin 277 | match eval (Put e1 e2) q with 278 | | Some (n, q') => ⟨c , n ,p, q', s⟩ 279 | | None => ⟪p, s⟫ 280 | end. 281 | = {auto} 282 | match eval e1 q with 283 | | Some (n, q') => match eval e2 n with 284 | | Some (n', q'') => ⟨c , n' ,p, q'', s⟩ 285 | | None => ⟪p, s⟫ 286 | end 287 | | None => ⟪p, s⟫ 288 | end. 289 | <|= {apply IHe2} 290 | match eval e1 q with 291 | | Some (n, q') => ⟨comp' e2 r c , n ,p, n, s⟩ 292 | | None => ⟪p, s⟫ 293 | end. 294 | <== {apply vm_write} 295 | match eval e1 q with 296 | | Some (n, q') => ⟨WRITE (comp' e2 r c) , n ,p, q', s⟩ 297 | | None => ⟪p, s⟫ 298 | end. 299 | <|= {apply IHe1} 300 | ⟨comp' e1 r (WRITE (comp' e2 r c)) , a ,p, q, s⟩. 301 | []. 302 | Qed. 303 | 304 | 305 | (** * Soundness *) 306 | 307 | (** Since the VM is defined as a small step operational semantics, we *) 308 | (* have to prove that the VM is deterministic and does not get stuck in *) 309 | (* order to derive soundness from the above theorem. *) 310 | 311 | 312 | Lemma determ_vm : determ VM. 313 | intros C C1 C2 V. induction V; intro V'; inversion V'; subst;eauto. 314 | rewrite H in *; dist. rewrite H in H3. dist. rewrite H in H6. dist. 315 | Qed. 316 | 317 | 318 | Theorem sound x a p q q' s n C : 319 | freeFrom first s -> eval x q = Some (n, q') 320 | -> ⟨comp x, a, p, q, s⟩ =>>! C -> exists s', C = ⟨HALT, n,p, q', s'⟩. 321 | Proof. 322 | intros F E M. 323 | pose (spec x first HALT a p s q F). unfold Reach in *. repeat autodestruct. 324 | pose (determ_trc determ_vm) as D. 325 | unfold determ in D. inversion H0. subst. eexists. eapply D. apply M. 326 | rewrite E in *. 327 | split. inversion H1. subst. 328 | 329 | apply H. intro Contra. destruct Contra. 330 | inversion H3. 331 | Qed. 332 | 333 | End LocalState. 334 | -------------------------------------------------------------------------------- /coq/WhileCoalesced.v: -------------------------------------------------------------------------------- 1 | (** Calculation for a language with state and while loops using a 2 | semantics with coalesced rules (one rule for each language 3 | construct). *) 4 | 5 | Require Import List. 6 | Require Import ListIndex. 7 | Require Import Tactics. 8 | Require Import Coq.Program.Equality. 9 | 10 | Require Import Init.Nat. 11 | 12 | 13 | Module While (Import mem : Memory). 14 | 15 | 16 | (** * Syntax *) 17 | 18 | Inductive Expr : Set := 19 | | Val : nat -> Expr 20 | | Add : Expr -> Expr -> Expr 21 | | Get : Expr 22 | | Put : Expr -> Expr -> Expr 23 | | While : Expr -> Expr -> Expr. 24 | 25 | Definition State := nat. 26 | Definition Value := nat. 27 | 28 | Reserved Notation "⟪ x , q ⟫ ⇓ ⟪ y , q' ⟫" (at level 80, no associativity). 29 | 30 | 31 | Inductive eval : Expr -> State -> Value -> State -> Prop := 32 | | eval_val n q : ⟪Val n, q⟫ ⇓ ⟪n, q⟫ 33 | | eval_add q q' q'' x y m n : ⟪x, q⟫ ⇓ ⟪m,q'⟫ -> ⟪y,q'⟫ ⇓ ⟪n, q''⟫ -> ⟪Add x y, q⟫ ⇓ ⟪m + n, q''⟫ 34 | | eval_get q : ⟪Get, q⟫ ⇓ ⟪q, q⟫ 35 | | eval_put q q' q'' v n x y : ⟪x, q⟫ ⇓ ⟪n,q'⟫ -> ⟪y, n⟫ ⇓ ⟪v, q''⟫ -> ⟪Put x y , q⟫ ⇓ ⟪v,q''⟫ 36 | | eval_while x y q n m q' q'' v q''': 37 | ⟪x, q⟫ ⇓ ⟪n,q'⟫ -> 38 | (n <> 0 -> ⟪y,q'⟫ ⇓ ⟪m,q''⟫) -> 39 | (n <> 0 -> ⟪While x y, q''⟫ ⇓ ⟪v,q'''⟫) -> 40 | (n = 0 -> v = 0) -> 41 | (n = 0 -> q''' = q') -> 42 | ⟪While x y, q⟫ ⇓ ⟪v,q'''⟫ 43 | where "⟪ x , q ⟫ ⇓ ⟪ y , q' ⟫" := (eval x q y q'). 44 | 45 | 46 | (** * Compiler *) 47 | 48 | Inductive Code : Set := 49 | | LOAD : nat -> Code -> Code 50 | | ADD : Reg -> Code -> Code 51 | | STORE : Reg -> Code -> Code 52 | | GET : Code -> Code 53 | | JUMP : Reg -> Code 54 | | JMPZ : Code -> Code -> Code 55 | | LABEL : Reg -> Code -> Code 56 | | PUT : Code -> Code 57 | | HALT : Code. 58 | 59 | Fixpoint comp' (e : Expr) (r : Reg) (c : Code) : Code := 60 | match e with 61 | | Val n => LOAD n c 62 | | Add x y => comp' x r (STORE r (comp' y (next r) (ADD r c))) 63 | | While x y => LABEL r (comp' x (next r) (JMPZ c (comp' y (next r) (JUMP r)))) 64 | | Get => GET c 65 | | Put x y => comp' x r (PUT (comp' y r c)) 66 | end. 67 | 68 | Definition comp (e : Expr) : Code := comp' e first HALT. 69 | 70 | 71 | Inductive Elem : Set := 72 | | VAL : Value -> Elem 73 | | CODE : Code -> Elem 74 | . 75 | 76 | 77 | 78 | (** * Virtual Machine *) 79 | 80 | Inductive Conf : Type := 81 | | conf : Code -> Value -> State -> Mem Elem -> Conf. 82 | 83 | Notation "⟨ c , a , q , s ⟩" := (conf c a q s). 84 | 85 | Reserved Notation "x ==> y" (at level 80, no associativity). 86 | Inductive VM : Conf -> Conf -> Prop := 87 | | vm_load n c s a q : ⟨LOAD n c, a, q, s⟩ ==> ⟨c, n, q, s⟩ 88 | | vm_add c m n r s q : s[r] = VAL m -> ⟨ADD r c, n, q, s⟩ 89 | ==> ⟨c, m + n, q, s⟩ 90 | | vm_store c v r s q : ⟨STORE r c, v, q, s⟩ 91 | ==> ⟨c, v, q, s[r:= VAL v]⟩ 92 | | vm_jump n r s q c : s[r] = CODE c -> ⟨JUMP r, n, q, s⟩ 93 | ==> ⟨c, n, q, s⟩ 94 | | vm_jmpz_zero s q c c' : ⟨JMPZ c' c, 0, q, s⟩ 95 | ==> ⟨c', 0, q, s⟩ 96 | | vm_jmpz_nzero n s q c c' : n <> 0 -> 97 | ⟨JMPZ c' c, n, q, s⟩ ==> ⟨c, n, q, s⟩ 98 | | vm_label n r s q c : ⟨LABEL r c, n, q, s⟩ ==> ⟨c, n, q, s[r:=CODE (LABEL r c)]⟩ 99 | | vm_get c s a q : ⟨GET c, a, q, s⟩ ==> ⟨c, q, q, s⟩ 100 | | vm_put c s a q : ⟨PUT c, a, q, s⟩ ==> ⟨c, a, a, s⟩ 101 | where "x ==> y" := (VM x y). 102 | 103 | (** Conversion functions from semantics to VM *) 104 | 105 | Inductive cle : Conf -> Conf -> Prop := 106 | | cle_mem c a e s s' : s ⊑ s' -> cle ⟨ c , a , e , s ⟩ ⟨ c , a , e , s' ⟩. 107 | 108 | Hint Constructors cle : core. 109 | 110 | Lemma nat_neq n m : false = (n =? m) -> n <> m. 111 | Proof. 112 | intro H. symmetry in H. rewrite <- PeanoNat.Nat.eqb_neq. assumption. 113 | Qed . 114 | 115 | Lemma nat_eq n m : true = (n =? m) -> n = m. 116 | Proof. 117 | intro H. symmetry in H. rewrite <- PeanoNat.Nat.eqb_eq. assumption. 118 | Qed . 119 | 120 | Hint Resolve nat_eq nat_neq : core. 121 | 122 | (** * Calculation *) 123 | 124 | (** Boilerplate to import calculation tactics *) 125 | Module VM <: Machine. 126 | Definition Conf := Conf. 127 | Definition Pre := cle. 128 | Definition Rel := VM. 129 | Definition MemElem := nat. 130 | Lemma monotone : monotonicity cle VM. 131 | prove_monotonicity. Qed. 132 | Lemma preorder : is_preorder cle. 133 | prove_preorder. Qed. 134 | End VM. 135 | Module VMCalc := Calculation mem VM. 136 | Import VMCalc. 137 | 138 | (** Specification of the compiler *) 139 | 140 | Theorem spec p v q q' r c a s : 141 | freeFrom r s -> ⟪p, q⟫ ⇓ ⟪v, q'⟫ -> 142 | ⟨comp' p r c, a, q, s⟩ =|> ⟨c , v, q', s⟩. 143 | 144 | (** Setup the induction proof *) 145 | 146 | Proof. 147 | intros F E. 148 | generalize dependent c. 149 | generalize dependent a. 150 | generalize dependent s. 151 | generalize dependent r. 152 | induction E; intros. 153 | 154 | (** Calculation of the compiler *) 155 | 156 | (** - [(Val n,q) ⇓ (n,q)]: *) 157 | 158 | begin 159 | ⟨c, n , q, s⟩. 160 | <== { apply vm_load } 161 | ⟨LOAD n c, a, q, s⟩. 162 | []. 163 | 164 | (** - [(Add x y,q) ⇓ (m + n,q')]: *) 165 | 166 | begin 167 | ⟨c, m + n, q'', s⟩. 168 | ⊑ { auto } 169 | ⟨c, m + n, q'', s[r:=VAL m]⟩ . 170 | <== { apply vm_add } 171 | ⟨ADD r c, n, q'', s[r:=VAL m]⟩ . 172 | <|= { apply IHE2 } 173 | ⟨comp' y (next r) (ADD r c), m, q', s[r:= VAL m]⟩ . 174 | <== { apply vm_store } 175 | ⟨STORE r (comp' y (next r) (ADD r c)), m, q', s⟩. 176 | <|= { apply IHE1 } 177 | ⟨comp' x r (STORE r (comp' y (next r) (ADD r c))), a, q, s⟩. 178 | []. 179 | 180 | (** - [(Get,q) ⇓ (q,q)]: *) 181 | 182 | begin 183 | ⟨ c, q, q, s ⟩. 184 | <== {apply vm_get} 185 | ⟨ GET c, a, q, s ⟩. 186 | []. 187 | 188 | (** - [(Put x y,q) ⇓ (v,q'')]: *) 189 | begin 190 | ⟨ c, v, q'', s ⟩. 191 | <|= { apply IHE2 } 192 | ⟨ comp' y r c, n, n, s ⟩. 193 | ⊑ {auto} 194 | ⟨ comp' y r c, n, n, s ⟩. 195 | <== { apply vm_put } 196 | ⟨ PUT (comp' y r c), n, q', s ⟩. 197 | <|= { apply IHE1 } 198 | ⟨ comp' x r (PUT (comp' y r c)), a, q, s ⟩. 199 | []. 200 | 201 | 202 | (** - [(While x y ,q) ⇓ (v,q''')]: *) 203 | begin 204 | ⟨ c, v, q''', s ⟩. 205 | <|= {apply H2} 206 | (if n =? 0 then ⟨ c, v, q''', s ⟩ 207 | else ⟨ comp' (While x y) r c, m, q'', s ⟩). 208 | <|= {rewrite H3, H4} 209 | (if n =? 0 then ⟨ c, 0, q', s ⟩ 210 | else ⟨ comp' (While x y) r c, m, q'', s ⟩). 211 | ⊑ { auto } 212 | (if n =? 0 then ⟨ c, 0, q', s ⟩ 213 | else ⟨ comp' (While x y) r c, m, q'', s[r:=CODE (comp' (While x y) r c)] ⟩). 214 | <== { apply vm_jump } 215 | (if n =? 0 then ⟨ c, 0, q', s ⟩ 216 | else ⟨ JUMP r, m, q'', s[r:=CODE (comp' (While x y) r c)] ⟩). 217 | <|= { apply H0 } 218 | (if n =? 0 then ⟨ c, 0, q', s ⟩ 219 | else ⟨ comp' y (next r) (JUMP r), n, q', s[r:=CODE (comp' (While x y) r c)] ⟩). 220 | <== { eapply vm_jmpz_nzero } 221 | (if n =? 0 then ⟨ c, 0, q', s ⟩ 222 | else ⟨ JMPZ c (comp' y (next r) (JUMP r)), n, q', s[r:=CODE (comp' (While x y) r c)] ⟩). 223 | <== { eapply vm_jmpz_zero} 224 | (if n =? 0 then ⟨ JMPZ c (comp' y (next r) (JUMP r)), 0, q', s ⟩ 225 | else ⟨ JMPZ c (comp' y (next r) (JUMP r)), n, q', s[r:=CODE (comp' (While x y) r c)] ⟩). 226 | ⊑ { auto } 227 | (if n =? 0 then ⟨ JMPZ c (comp' y (next r) (JUMP r)), 0, q', s[r:=CODE (comp' (While x y) r c)] ⟩ 228 | else ⟨ JMPZ c (comp' y (next r) (JUMP r)), n, q', s[r:=CODE (comp' (While x y) r c)] ⟩). 229 | <|= { (assert (n = 0) by auto; subst) } 230 | ⟨ JMPZ c (comp' y (next r) (JUMP r)), n, q', s[r:=CODE (comp' (While x y) r c)] ⟩. 231 | <|= { apply IHE } 232 | ⟨ comp' x (next r) (JMPZ c (comp' y (next r) (JUMP r))), a, q, s[r:=CODE (comp' (While x y) r c)] ⟩. 233 | <== { apply vm_label } 234 | ⟨ LABEL r (comp' x (next r) (JMPZ c (comp' y (next r) (JUMP r)))), a, q, s ⟩. 235 | []. 236 | Qed. 237 | 238 | (** * Soundness *) 239 | 240 | Lemma determ_vm : determ VM. 241 | intros C C1 C2 V. induction V; intro V'; inversion V'; subst; congruence. 242 | Qed. 243 | 244 | 245 | Definition terminates (p : Expr) : Prop := exists n q, ⟪p, 0⟫ ⇓ ⟪n,q⟫. 246 | 247 | Theorem sound p a s C : freeFrom first s -> terminates p -> ⟨comp p, a, 0, s⟩ =>>! C -> 248 | exists n s' q, C = ⟨HALT , n, q, s'⟩ /\ ⟪p,0⟫ ⇓ ⟪n,q⟫. 249 | Proof. 250 | unfold terminates. intros F T M. destruct T as [n T]. destruct T as [q T]. 251 | pose (spec p n 0 q first HALT a s F T) as H'. 252 | unfold Reach in *. repeat autodestruct. 253 | pose (determ_trc determ_vm) as D. 254 | unfold determ in D. inversion H0. subst. exists n. eexists. exists q. split. eapply D. apply M. split. 255 | unfold comp. 256 | simpl in *. apply H. intro Contra. destruct Contra. 257 | inversion H1. assumption. 258 | Qed. 259 | 260 | 261 | (* Finally, we show that the semantics of Expr is equivalent to the 262 | semantics below, which is in a more conventional style. *) 263 | 264 | Reserved Notation "⟪ x , q ⟫ ⇓' ⟪ y , q' ⟫" (at level 80, no associativity). 265 | 266 | Inductive eval' : Expr -> State -> Value -> State -> Prop := 267 | | eval_val' n q : ⟪Val n, q⟫ ⇓' ⟪n, q⟫ 268 | | eval_add' q q' q'' x y m n : 269 | ⟪x, q⟫ ⇓' ⟪m,q'⟫ -> ⟪y,q'⟫ ⇓' ⟪n, q''⟫ -> ⟪Add x y, q⟫ ⇓' ⟪m + n, q''⟫ 270 | | eval_get' q : ⟪Get, q⟫ ⇓' ⟪q, q⟫ 271 | | eval_put' q q' q'' v n x y : 272 | ⟪x, q⟫ ⇓' ⟪n,q'⟫ -> ⟪y, n⟫ ⇓' ⟪v, q''⟫ -> ⟪Put x y , q⟫ ⇓' ⟪v,q''⟫ 273 | | eval_step' x y q n m q' q'' v q''': 274 | ⟪x, q⟫ ⇓' ⟪n,q'⟫ -> 275 | n <> 0 -> ⟪y,q'⟫ ⇓' ⟪m,q''⟫ -> ⟪While x y, q''⟫ ⇓' ⟪v,q'''⟫ -> 276 | ⟪While x y, q⟫ ⇓' ⟪v,q'''⟫ 277 | | eval_while_stop' x y q n q' : 278 | ⟪x, q⟫ ⇓' ⟪n,q'⟫ -> 279 | n = 0 -> 280 | ⟪While x y, q⟫ ⇓' ⟪0,q'⟫ 281 | where "⟪ x , q ⟫ ⇓' ⟪ y , q' ⟫" := (eval' x q y q'). 282 | 283 | Hint Constructors eval eval' : core. 284 | 285 | Theorem sem_eq x q y q' : ⟪ x , q ⟫ ⇓' ⟪ y , q' ⟫ <-> ⟪ x , q ⟫ ⇓ ⟪ y , q' ⟫. 286 | Proof. 287 | split;intro E. 288 | * induction E;eauto. 289 | - econstructor; intros; try first [eassumption|contradiction]. 290 | - econstructor; intros; try first [eauto|contradiction]; subst; contradiction H0; auto. 291 | Unshelve. auto. auto. 292 | * induction E;eauto. 293 | remember (n =? 0) as Q. destruct Q. 294 | - apply nat_eq in HeqQ. rewrite H3, H4 by assumption. eapply eval_while_stop'. 295 | apply IHE. assumption. 296 | - apply nat_neq in HeqQ. eauto. 297 | Qed. 298 | 299 | End While. 300 | -------------------------------------------------------------------------------- /coq/State.v: -------------------------------------------------------------------------------- 1 | (** Calculation of the simple arithmetic language with exceptions and 2 | (global) state. *) 3 | 4 | Require Import List. 5 | Require Import Tactics. 6 | Require Import Coq.Program.Equality. 7 | Module State (Import mem : Memory). 8 | 9 | 10 | (** * Syntax *) 11 | 12 | Inductive Expr : Set := 13 | | Val : nat -> Expr 14 | | Add : Expr -> Expr -> Expr 15 | | Throw : Expr 16 | | Catch : Expr -> Expr -> Expr 17 | | Get : Expr 18 | | Put : Expr -> Expr -> Expr. 19 | 20 | Definition state := nat. 21 | 22 | (** * Semantics *) 23 | 24 | Fixpoint eval (x: Expr) (q : state) : option nat * state := 25 | match x with 26 | | Val n => (Some n, q) 27 | | Add x1 x2 => match eval x1 q with 28 | | (Some m, q') => match eval x2 q' with 29 | | (Some n, q'') => (Some (m + n), q'') 30 | | (None, q'') => (None, q'') 31 | end 32 | | (None, q') => (None, q') 33 | end 34 | | Throw => (None, q) 35 | | Catch x h => match eval x q with 36 | | (Some n, q') => (Some n, q') 37 | | (None, q') => eval h q' 38 | end 39 | | Get => (Some q, q) 40 | | Put x y => match eval x q with 41 | | (Some n, q') => eval y n 42 | | (None, q') => (None, q') 43 | end 44 | end. 45 | 46 | (** * Compiler *) 47 | 48 | Inductive Code : Set := 49 | | LOAD : nat -> Code -> Code 50 | | ADD : Reg -> Code -> Code 51 | | STORE : Reg -> Code -> Code 52 | | THROW : Code 53 | | UNMARK : Code -> Code 54 | | MARK : Reg -> Code -> Code -> Code 55 | | READ : Code -> Code 56 | | WRITE : Code -> Code 57 | | HALT : Code. 58 | 59 | Fixpoint comp' (x : Expr) (r : Reg) (c : Code) : Code := 60 | match x with 61 | | Val n => LOAD n c 62 | | Add e1 e2 => comp' e1 r (STORE r (comp' e2 (next r) (ADD r c))) 63 | | Throw => THROW 64 | | Catch e1 e2 => MARK r (comp' e2 r c) (comp' e1 (next r) (UNMARK c)) 65 | | Get => READ c 66 | | Put x y => comp' x r (WRITE (comp' y r c)) 67 | end. 68 | 69 | Definition comp (x : Expr) : Code := comp' x first HALT. 70 | 71 | Definition Han := option Reg. 72 | 73 | Inductive RVal : Set := 74 | | VAL : nat -> RVal 75 | | HAN : Code -> Han -> RVal. 76 | 77 | (** * Virtual Machine *) 78 | 79 | Inductive Conf' : Type := 80 | | conf : Code -> nat -> Han -> state -> Conf' 81 | | fail : Han -> state -> Conf'. 82 | 83 | Definition Conf : Type := Conf' * Mem RVal. 84 | 85 | 86 | Notation "⟨ x , y , z , q , s ⟩" := (conf x y z q, s). 87 | Notation "⟪ x , q , s ⟫" := (fail x q, s). 88 | 89 | Reserved Notation "x ==> y" (at level 80, no associativity). 90 | Inductive VM : Conf -> Conf -> Prop := 91 | | vm_load n a c s q p : ⟨LOAD n c, a , p, q, s⟩ ==> ⟨c , n, p, q, s⟩ 92 | | vm_add c s a r n q p : s[r]= VAL n -> ⟨ADD r c, a , p, q, s⟩ ==> ⟨c , n + a, p, q, s⟩ 93 | | vm_store c s a r q p : ⟨STORE r c, a, p, q, s⟩ ==> ⟨c , a, p, q, s[r:=VAL a]⟩ 94 | | vm_throw a s q p : ⟨THROW, a, p, q, s⟩ ==> ⟪p, q, s⟫ 95 | | vm_fail p p' s q c : s[p] = HAN c p' -> ⟪ Some p, q, s⟫ ==> ⟨c, 0, p', q, s⟩ 96 | | vm_unmark p p' s a q c c' : s[p] = HAN c' p' -> 97 | ⟨UNMARK c, a, Some p, q, s⟩ ==> ⟨c, a, p', q, s⟩ 98 | | vm_mark p r s a c c' q : ⟨MARK r c' c, a, p, q, s⟩ ==> ⟨c, a, Some r, q, s[r:= HAN c' p]⟩ 99 | | vm_read a c s q p : ⟨READ c, a , p, q, s⟩ ==> ⟨c , q, p, q, s⟩ 100 | | vm_write a c s q p : ⟨WRITE c, a , p, q, s⟩ ==> ⟨c , a, p, a, s⟩ 101 | where "x ==> y" := (VM x y). 102 | 103 | Inductive cle : Conf -> Conf -> Prop := 104 | | cle_mem f s s' : s ⊑ s' -> cle (f, s) (f, s'). 105 | 106 | Hint Constructors cle : core. 107 | 108 | 109 | (** * Calculation *) 110 | 111 | (** Boilerplate to import calculation tactics *) 112 | Module VM <: Machine. 113 | Definition Conf := Conf. 114 | Definition Pre := cle. 115 | Definition Rel := VM. 116 | Lemma monotone : monotonicity cle VM. 117 | prove_monotonicity. Qed. 118 | Lemma preorder : is_preorder cle. 119 | prove_preorder. Qed. 120 | End VM. 121 | 122 | Module VMCalc := Calculation mem VM. 123 | Import VMCalc. 124 | 125 | 126 | (** Specification of the compiler *) 127 | 128 | Theorem spec : forall e r c a p s q, 129 | freeFrom r s -> 130 | ⟨comp' e r c, a, p, q, s⟩ 131 | =|> 132 | match eval e q with 133 | | (Some n, q') => ⟨c , n, p, q', s⟩ 134 | | (None, q') => ⟪p, q', s⟫ 135 | end. 136 | 137 | 138 | (** Setup the induction proof *) 139 | 140 | Proof. 141 | induction e;intros. 142 | 143 | (** Calculation of the compiler *) 144 | 145 | (** - [x = Val n]: *) 146 | 147 | begin 148 | match eval (Val n) q with 149 | | (Some n', q') => ⟨c , n' ,p, q', s⟩ 150 | | (None, q') => ⟪p, q', s⟫ 151 | end. 152 | = {auto} 153 | ⟨ c, n, p, q, s⟩. 154 | <== {apply vm_load} 155 | ⟨ LOAD n c, a, p, q, s ⟩. 156 | []. 157 | 158 | (** - [x = Add x1 x2]: *) 159 | 160 | begin 161 | match eval (Add e1 e2) q with 162 | | (Some n, q') => ⟨c , n ,p, q', s⟩ 163 | | (None, q') => ⟪p, q', s⟫ 164 | end. 165 | = {auto} 166 | match eval e1 q with 167 | | (Some n, q') => match eval e2 q' with 168 | | (Some n', q'') => ⟨c , n + n' , p, q'', s⟩ 169 | | (None, q'') => ⟪p, q'', s⟫ 170 | end 171 | | (None, q') => ⟪p, q', s⟫ 172 | end. 173 | ⊑ { auto } 174 | match eval e1 q with 175 | | (Some n, q') => match eval e2 q' with 176 | | (Some n', q'') => ⟨c , n + n' ,p, q'', s[r:=VAL n]⟩ 177 | | (None, q'') => ⟪p, q'', s[r:=VAL n]⟫ 178 | end 179 | | (None, q') => ⟪p, q', s⟫ 180 | end. 181 | <== { apply vm_add } 182 | match eval e1 q with 183 | | (Some n, q') => match eval e2 q' with 184 | | (Some n', q'') => ⟨ADD r c , n' ,p, q'', s[r:=VAL n]⟩ 185 | | (None, q'') => ⟪p, q'', s[r:=VAL n]⟫ 186 | end 187 | | (None, q') => ⟪p, q', s⟫ 188 | end. 189 | <|= {apply IHe2} 190 | match eval e1 q with 191 | | (Some n, q') => ⟨comp' e2 (next r) (ADD r c) , n ,p, q', s[r:=VAL n]⟩ 192 | | (None, q') => ⟪p, q', s⟫ 193 | end. 194 | <== { apply vm_store } 195 | match eval e1 q with 196 | | (Some n, q') => ⟨STORE r (comp' e2 (next r) (ADD r c)) , n ,p, q', s⟩ 197 | | (None, q') => ⟪p, q', s⟫ 198 | end. 199 | <|= { apply IHe1 } 200 | ⟨comp' e1 r (STORE r (comp' e2 (next r) (ADD r c))), a,p, q, s⟩. 201 | []. 202 | 203 | (** - [x = Throw]: *) 204 | 205 | begin 206 | match eval Throw q with 207 | | (Some n, q') => ⟨c , n ,p, q', s⟩ 208 | | (None, q') => ⟪p, q', s⟫ 209 | end. 210 | = { auto } 211 | ⟪p, q, s⟫. 212 | <== { apply vm_throw } 213 | ⟨THROW, a, p, q, s⟩. 214 | []. 215 | 216 | (** - [x = Catch x1 x2]: *) 217 | 218 | begin 219 | match eval (Catch e1 e2) q with 220 | | (Some n, q') => ⟨c , n ,p, q', s⟩ 221 | | (None, q') => ⟪p, q', s⟫ 222 | end. 223 | = { auto } 224 | match eval e1 q with 225 | | (Some n, q') => ⟨c , n ,p, q', s⟩ 226 | | (None, q') => match eval e2 q' with 227 | | (Some n', q'') => ⟨c , n' ,p, q'', s⟩ 228 | | (None, q'') => ⟪p, q'', s⟫ 229 | end 230 | end. 231 | <|= { apply IHe2 } 232 | match eval e1 q with 233 | | (Some n, q') => ⟨c , n ,p, q', s⟩ 234 | | (None, q') => ⟨comp' e2 r c , 0 ,p, q', s⟩ 235 | end. 236 | ⊑ { eauto } 237 | match eval e1 q with 238 | | (Some n, q') => ⟨c , n ,p, q', s⟩ 239 | | (None, q') => ⟨comp' e2 r c , 0 ,p, q', s[r:= HAN (comp' e2 r c) p]⟩ 240 | end. 241 | <== {apply vm_fail} 242 | match eval e1 q with 243 | | (Some n, q') => ⟨c , n ,p, q', s⟩ 244 | | (None, q') => ⟪Some r, q', s[r:= HAN (comp' e2 r c) p]⟫ 245 | end. 246 | ⊑ {auto} 247 | match eval e1 q with 248 | | (Some n, q') => ⟨c , n ,p, q', s[r:= HAN (comp' e2 r c) p]⟩ 249 | | (None, q') => ⟪Some r, q', s[r:= HAN (comp' e2 r c) p]⟫ 250 | end. 251 | <== {eapply vm_unmark} 252 | match eval e1 q with 253 | | (Some n, q') => ⟨UNMARK c , n,Some r, q', s[r:= HAN (comp' e2 r c) p]⟩ 254 | | (None, q') => ⟪Some r, q', s[r:= HAN (comp' e2 r c) p]⟫ 255 | end. 256 | <|= {apply IHe1} 257 | ⟨comp' e1 (next r) (UNMARK c) , a ,Some r, q, s[r:= HAN (comp' e2 r c) p]⟩. 258 | <== {apply vm_mark} 259 | ⟨MARK r (comp' e2 r c) (comp' e1 (next r) (UNMARK c)) , a ,p, q, s⟩. 260 | []. 261 | 262 | 263 | begin 264 | match eval Get q with 265 | | (Some n, q') => ⟨c , n ,p, q', s⟩ 266 | | (None, q') => ⟪p, q', s⟫ 267 | end. 268 | = {auto} 269 | ⟨c , q ,p, q, s⟩. 270 | <== {apply vm_read} 271 | ⟨READ c , a ,p, q, s⟩. 272 | []. 273 | 274 | begin 275 | match eval (Put e1 e2) q with 276 | | (Some n, q') => ⟨c , n ,p, q', s⟩ 277 | | (None, q') => ⟪p, q', s⟫ 278 | end. 279 | = {auto} 280 | match eval e1 q with 281 | | (Some n, q') => match eval e2 n with 282 | | (Some n', q'') => ⟨c , n' ,p, q'', s⟩ 283 | | (None, q'') => ⟪p, q'', s⟫ 284 | end 285 | | (None, q') => ⟪p, q', s⟫ 286 | end. 287 | <|= {apply IHe2} 288 | match eval e1 q with 289 | | (Some n, q') => ⟨comp' e2 r c , n ,p, n, s⟩ 290 | | (None, q') => ⟪p, q', s⟫ 291 | end. 292 | <== {apply vm_write} 293 | match eval e1 q with 294 | | (Some n, q') => ⟨WRITE (comp' e2 r c) , n ,p, q', s⟩ 295 | | (None, q') => ⟪p, q', s⟫ 296 | end. 297 | <|= {apply IHe1} 298 | ⟨comp' e1 r (WRITE (comp' e2 r c)) , a ,p, q, s⟩. 299 | []. 300 | Qed. 301 | 302 | 303 | (** * Soundness *) 304 | 305 | (** Since the VM is defined as a small step operational semantics, we *) 306 | (* have to prove that the VM is deterministic and does not get stuck in *) 307 | (* order to derive soundness from the above theorem. *) 308 | 309 | 310 | Lemma determ_vm : determ VM. 311 | intros C C1 C2 V. induction V; intro V'; inversion V'; subst;eauto. 312 | rewrite H in *; dist. rewrite H in H4. dist. rewrite H in H6. dist. 313 | Qed. 314 | 315 | 316 | Theorem sound x a p q q' s n C : 317 | freeFrom first s -> eval x q = (Some n, q') 318 | -> ⟨comp x, a, p, q, s⟩ =>>! C -> exists s', C = ⟨HALT, n,p, q', s'⟩. 319 | Proof. 320 | intros F E M. 321 | pose (spec x first HALT a p s q F). unfold Reach in *. repeat autodestruct. 322 | pose (determ_trc determ_vm) as D. 323 | unfold determ in D. inversion H0. subst. eexists. eapply D. apply M. 324 | rewrite E in *. 325 | split. inversion H1. subst. 326 | 327 | apply H. intro Contra. destruct Contra. 328 | inversion H3. 329 | Qed. 330 | 331 | End State. 332 | -------------------------------------------------------------------------------- /coq/LambdaException.v: -------------------------------------------------------------------------------- 1 | (** Calculation for the lambda calculus + arithmetic + exceptions. *) 2 | 3 | Require Import List. 4 | Require Import ListIndex. 5 | Require Import Tactics. 6 | Require Import Coq.Program.Equality. 7 | Module Lambda (Import mem : TruncMem). 8 | 9 | 10 | (** * Syntax *) 11 | 12 | Inductive Expr : Set := 13 | | Val : nat -> Expr 14 | | Add : Expr -> Expr -> Expr 15 | | Var : nat -> Expr 16 | | Abs : Expr -> Expr 17 | | App : Expr -> Expr -> Expr 18 | | Throw : Expr 19 | | Catch : Expr -> Expr -> Expr. 20 | 21 | (** * Semantics *) 22 | 23 | (** The evaluator for this language is given as follows (as in the 24 | paper): 25 | << 26 | type Env = [Value] 27 | data Value = Num Int | Fun (Value -> Maybe Value) 28 | 29 | 30 | eval :: Expr -> Env -> Maybe Value 31 | eval (Val n) e = Some (Num n) 32 | eval (Add x y) e = case eval x e of 33 | Just (Num n) -> case eval y e of 34 | Just (Num m) -> Num (n+m) 35 | _ -> Nothing 36 | _ -> Nothing 37 | eval (Var i) e = if i < length e then Just (e !! i) else Nothing 38 | eval (Abs x) e = Just (Fun (\v -> eval x (v:e))) 39 | eval (App x y) e = case eval x e of 40 | Just (Fun f) -> eval y e >>= f 41 | _ -> Nothing 42 | 43 | >> 44 | After defunctionalisation and translation into relational form we 45 | obtain the semantics below. *) 46 | 47 | Inductive Value : Set := 48 | | Num : nat -> Value 49 | | Clo : Expr -> list Value -> Value. 50 | 51 | Definition Env := list Value. 52 | 53 | Reserved Notation "x ⇓[ e ] y" (at level 80, no associativity). 54 | 55 | Inductive eval : Expr -> Env -> option Value -> Prop := 56 | | eval_val e n : Val n ⇓[e] Some (Num n) 57 | | eval_add e x y m n : x ⇓[e] m -> y ⇓[e] n 58 | -> Add x y ⇓[e] (match m, n with 59 | | Some (Num m'), Some (Num n') => Some (Num (m' + n')) 60 | | _,_ => None 61 | end ) 62 | | eval_throw e : Throw ⇓[e] None 63 | | eval_catch e x y m n : x ⇓[e] m -> y ⇓[e] n 64 | -> Catch x y ⇓[e] (match m with 65 | | None => n 66 | | _ => m 67 | end ) 68 | | eval_var e i : Var i ⇓[e] nth e i 69 | | eval_abs e x : Abs x ⇓[e] Some (Clo x e) 70 | | eval_app e x x'' y x' e' y' : x ⇓[e] Some (Clo x' e') -> y ⇓[e] Some y' -> x' ⇓[y' :: e'] x'' 71 | -> App x y ⇓[e] x'' 72 | | eval_app_fail x x' y y' e : x ⇓[e] x' -> y ⇓[e] y' -> 73 | (x' = None \/ exists n, x' = Some (Num n) \/ y' = None) -> 74 | App x y ⇓[e] None 75 | where "x ⇓[ e ] y" := (eval x e y). 76 | 77 | (** * Compiler *) 78 | 79 | Inductive Code : Set := 80 | | LOAD : nat -> Code -> Code 81 | | ADD : Reg -> Code -> Code 82 | | STORE : Reg -> Code -> Code 83 | | STC : Reg -> Code -> Code 84 | | LOOKUP : nat -> Code -> Code 85 | | RET : Code 86 | | APP : Reg -> Code -> Code 87 | | ABS : Code -> Code -> Code 88 | | THROW : Code 89 | | UNMARK : Code -> Code 90 | | MARK : Reg -> Code -> Code -> Code 91 | | HALT : Code. 92 | 93 | Fixpoint comp (e : Expr) (r : Reg) (c : Code) : Code := 94 | match e with 95 | | Val n => LOAD n c 96 | | Add x y => comp x r (STORE r (comp y (next r) (ADD r c))) 97 | | Var i => LOOKUP i c 98 | | App x y => comp x r (STC r (comp y (next r) (APP r c))) 99 | | Abs x => ABS (comp x (next first) RET) c 100 | | Throw => THROW 101 | | Catch e1 e2 => MARK r (comp e2 r c) (comp e1 (next r) (UNMARK c)) 102 | end. 103 | 104 | Definition compile (e : Expr) : Code := comp e first HALT. 105 | 106 | (** * Virtual Machine *) 107 | 108 | Inductive Value' : Set := 109 | | Num' : nat -> Value' 110 | | Clo' : Code -> list Value' -> Value' 111 | | Exc' : Value'. 112 | 113 | Definition Env' := list Value'. 114 | 115 | Definition Han := option Reg. 116 | 117 | 118 | Inductive RVal : Set := 119 | | VAL : nat -> RVal 120 | | CLO : Code -> Env' -> RVal 121 | | HAN : Code -> Env' -> Han -> RVal 122 | . 123 | 124 | Inductive SElem := 125 | | MEM : Mem RVal -> SElem 126 | | MARKED : SElem. 127 | 128 | Definition Lam : Type := list SElem. 129 | 130 | Inductive Conf' : Type := 131 | | conf : Code -> Value' -> Env' -> Han -> Conf' 132 | | fail : Han -> Conf'. 133 | 134 | Definition Conf : Type := Conf' * Lam * Mem RVal. 135 | 136 | Notation "⟨ x , y , z , w , k , s ⟩" := (conf x y z w, k, s). 137 | Notation "⟪ x , k , s ⟫" := (fail x, k, s). 138 | 139 | Definition empty := (@empty RVal). 140 | 141 | Reserved Notation "x ==> y" (at level 80, no associativity). 142 | Inductive VM : Conf -> Conf -> Prop := 143 | | vm_load n c s a e k h : ⟨LOAD n c, a, e, h, k, s⟩ ==> ⟨c, Num' n, e, h, k, s⟩ 144 | | vm_add c m n r s e k h : s[r] = VAL m -> ⟨ADD r c, Num' n, e, h, k, s⟩ 145 | ==> ⟨c, Num'(m + n), e, h, k, s⟩ 146 | | vm_add_fail s c c' e e' h r k : ⟨ADD r c, Clo' c' e', e, h, k, s⟩ ==> ⟪h, k, s⟫ 147 | | vm_store c n r s e k h : ⟨STORE r c, Num' n, e, h, k, s⟩ 148 | ==> ⟨c, Num' n, e, h, k, s[r:=VAL n]⟩ 149 | | vm_store_fail c r s e k h c' e' : ⟨STORE r c, Clo' c' e', e, h, k, s⟩ 150 | ==> ⟪h, k, s⟫ 151 | | vm_stc c c' e' r s e k h : ⟨STC r c, Clo' c' e', e, h, k, s⟩ 152 | ==> ⟨c, Clo' c' e', e, h, k, s[r:=CLO c' e']⟩ 153 | | vm_stc_fail c n r s e k h : ⟨STC r c, Num' n, e, h, k, s⟩ 154 | ==> ⟪h, k, s⟫ 155 | | vm_lookup e i c v a s k h : nth e i = Some v -> ⟨LOOKUP i c, a, e, h, k, s⟩ ==> ⟨c, v, e, h, k, s⟩ 156 | | vm_lookup_fail e i c a s k h : nth e i = None -> ⟨LOOKUP i c, a, e, h, k, s⟩ ==> ⟪h, k, s⟫ 157 | | vm_ret a c e e' s s' k h : s[first] = CLO c e -> ⟨RET, a, e', h, MEM s' :: k, s⟩ ==> ⟨c, a, e, h, k, s'⟩ 158 | | vm_app c c' e e' v r s k h : 159 | s[r]=CLO c' e' -> 160 | ⟨APP r c, v, e, h, k,s⟩ ==> ⟨c', Num' 0, v :: e', h, MEM (truncate r s) :: k, empty[first:=CLO c e]⟩ 161 | | vm_abs a c c' s e k h : ⟨ABS c' c, a, e, h, k, s⟩ ==> ⟨c, Clo' c' e, e, h, k, s⟩ 162 | | vm_throw a e h k s : ⟨THROW, a, e, h, k, s⟩ ==> ⟪h, k, s⟫ 163 | | vm_fail p p' s c k e : s[p] = HAN c e p' -> ⟪ Some p, MARKED :: k, s⟫ ==> ⟨c, Num' 0, e, p', k, s⟩ 164 | | vm_fail_pop p s s' k : ⟪ p, MEM s' :: k, s⟫ ==> ⟪ p, k, s'⟫ 165 | | vm_unmark p p' s a c c' e e' k : s[p] = HAN c' e' p' -> ⟨UNMARK c, a, e, Some p, MARKED :: k, s⟩ ==> ⟨c, a, e, p', k, s⟩ 166 | | vm_mark p r s a c c' e k : ⟨MARK r c' c, a, e, p, k, s⟩ ==> ⟨c, a, e, Some r, MARKED :: k, s[r:= HAN c' e p]⟩ 167 | where "x ==> y" := (VM x y). 168 | 169 | (** Conversion functions from semantics to VM *) 170 | 171 | Fixpoint conv (v : Value) : Value' := 172 | match v with 173 | | Num n => Num' n 174 | | Clo x e => Clo' (comp x (next first) RET) (map conv e) 175 | end. 176 | 177 | Definition convE : Env -> Env' := map conv. 178 | 179 | Inductive stackle : Lam -> Lam -> Prop := 180 | | stackle_empty : stackle nil nil 181 | | stackle_cons_mem s s' k k' : s ⊑ s' -> stackle k k' -> stackle (MEM s :: k) (MEM s' :: k') 182 | | stackle_cons_han k k' : stackle k k' -> stackle (MARKED :: k) (MARKED :: k'). 183 | 184 | Hint Constructors stackle : memory. 185 | 186 | Lemma stackle_refl k : stackle k k. 187 | Proof. 188 | induction k; try destruct a; constructor; auto with memory. 189 | Qed. 190 | 191 | Lemma stackle_trans k1 k2 k3 : stackle k1 k2 -> stackle k2 k3 -> stackle k1 k3. 192 | Proof. 193 | intros L1. generalize k3. induction L1; intros k3' L2; solve[assumption| inversion L2; subst; constructor; eauto with memory]. 194 | Qed. 195 | 196 | Hint Resolve stackle_refl stackle_trans : core. 197 | 198 | Inductive cle : Conf -> Conf -> Prop := 199 | | cle_mem f k k' s s' : stackle k k' -> s ⊑ s' -> cle (f, k, s) (f , k', s' ). 200 | 201 | Hint Constructors cle : core. 202 | 203 | Lemma rel_eq {T} {R : T -> T -> Prop} x y y' : R x y' -> y = y' -> R x y. 204 | Proof. intros. subst. auto. 205 | Qed . 206 | Lemma rel_eq' {T} {R : T -> T -> Prop} x x' y : R x' y -> x = x' -> R x y. 207 | Proof. intros. subst. auto. 208 | Qed . 209 | 210 | Ltac apply_eq t := eapply rel_eq; [apply t | repeat rewrite set_set; auto]. 211 | 212 | 213 | (** * Calculation *) 214 | 215 | (** Boilerplate to import calculation tactics *) 216 | 217 | Module VM <: Machine. 218 | Definition Conf := Conf. 219 | Definition Pre := cle. 220 | Definition Rel := VM. 221 | 222 | Lemma monotone : monotonicity cle VM. 223 | prove_monotonicity1; 224 | try (match goal with [H : stackle (_ :: _) _ |- _] => inversion H end) 225 | ; prove_monotonicity2. 226 | Qed. 227 | 228 | Lemma preorder : is_preorder cle. 229 | prove_preorder. Qed. 230 | End VM. 231 | 232 | 233 | 234 | Module VMCalc := Calculation mem VM. 235 | Import VMCalc. 236 | 237 | 238 | 239 | (** Specification of the compiler *) 240 | 241 | Theorem spec p v e r c a s k h : 242 | freeFrom r s -> p ⇓[e] v -> 243 | ⟨comp p r c, a, convE e, h, k, s⟩ =|> match v with 244 | | None => ⟪ h , k , s ⟫ 245 | | Some v' => ⟨c , conv v', convE e, h, k, s⟩ 246 | end. 247 | 248 | (** Setup the induction proof *) 249 | 250 | Proof. 251 | intros F E. 252 | generalize dependent c. 253 | generalize dependent a. 254 | generalize dependent s. 255 | generalize dependent r. 256 | generalize dependent k. 257 | generalize dependent h. 258 | induction E;intros. 259 | 260 | (** Calculation of the compiler *) 261 | 262 | (** - [Val n ⇓[e] Num n]: *) 263 | 264 | begin 265 | ⟨c, Num' n , convE e, h, k, s⟩. 266 | <== { apply vm_load } 267 | ⟨LOAD n c, a, convE e, h, k, s⟩. 268 | []. 269 | 270 | (** - [Add x y ⇓[e] Num (m + n)]: *) 271 | 272 | begin 273 | match m with 274 | | Some (Num m') => match n with 275 | | Some (Num n') => ⟨ c, Num' (m' + n'), convE e, h, k, s ⟩ 276 | | _ => ⟪ h, k, s ⟫ 277 | end 278 | | _ => ⟪ h, k, s ⟫ 279 | end. 280 | ⊑ {auto} 281 | match m with 282 | | Some (Num m') => match n with 283 | | Some (Num n') => ⟨ c, Num' (m' + n'), convE e, h, k, s[r:=VAL m'] ⟩ 284 | | _ => ⟪ h, k, s[r:=VAL m'] ⟫ 285 | end 286 | | _ => ⟪ h, k, s ⟫ 287 | end. 288 | <== { apply vm_add } 289 | match m with 290 | | Some (Num m') => match n with 291 | | Some (Num n') => ⟨ADD r c, Num' n', convE e, h, k, s[r:=VAL m'] ⟩ 292 | | _ => ⟪ h, k, s[r:=VAL m'] ⟫ 293 | end 294 | | _ => ⟪ h, k, s ⟫ 295 | end. 296 | <== { apply vm_add_fail } 297 | match m with 298 | | Some (Num m') => match n with 299 | | Some v => ⟨ADD r c, conv v, convE e, h, k, s[r:=VAL m'] ⟩ 300 | | None => ⟪ h, k, s[r:=VAL m'] ⟫ 301 | end 302 | | _ => ⟪ h, k, s ⟫ 303 | end. 304 | <|= { apply IHE2 } 305 | match m with 306 | | Some (Num m') => ⟨comp y (next r) (ADD r c), Num' m', convE e, h, k, s[r:=VAL m'] ⟩ 307 | | _ => ⟪ h, k, s ⟫ 308 | end. 309 | <== { apply vm_store } 310 | match m with 311 | | Some (Num m') => ⟨STORE r (comp y (next r) (ADD r c)), Num' m', convE e, h, k, s ⟩ 312 | | _ => ⟪ h, k, s ⟫ 313 | end. 314 | <== { apply vm_store_fail } 315 | match m with 316 | | Some v => ⟨STORE r (comp y (next r) (ADD r c)), conv v, convE e, h, k, s ⟩ 317 | | _ => ⟪ h, k, s ⟫ 318 | end. 319 | <|= { apply IHE1 } 320 | ⟨comp x r (STORE r (comp y (next r) (ADD r c))), a, convE e, h, k, s⟩. 321 | []. 322 | 323 | (** - [Throw ⇓[e] None] *) 324 | 325 | begin 326 | ⟪h, k, s⟫. 327 | <== {apply vm_throw} 328 | ⟨ THROW, a, convE e, h, k, s⟩. 329 | []. 330 | 331 | 332 | begin 333 | match m with 334 | | Some v'' => ⟨ c, conv v'', convE e, h, k, s ⟩ 335 | | None => match n with 336 | | Some v'' => ⟨ c, conv v'', convE e, h, k, s ⟩ 337 | | None => ⟪ h, k, s ⟫ 338 | end 339 | end. 340 | <|= {apply IHE2} 341 | match m with 342 | | Some v'' => ⟨ c, conv v'', convE e, h, k, s ⟩ 343 | | None => ⟨ comp y r c, Num' 0, convE e, h, k, s ⟩ 344 | end. 345 | ⊑ {eauto} 346 | match m with 347 | | Some v'' => ⟨ c, conv v'', convE e, h, k, s ⟩ 348 | | None => ⟨ comp y r c, Num' 0, convE e, h, k, s[r:= HAN (comp y r c) (convE e) h] ⟩ 349 | end. 350 | <== {apply vm_fail} 351 | match m with 352 | | Some v'' => ⟨ c, conv v'', convE e, h, k, s ⟩ 353 | | None => ⟪ Some r, MARKED :: k, s[r:= HAN (comp y r c) (convE e) h] ⟫ 354 | end. 355 | ⊑ {eauto} 356 | match m with 357 | | Some v'' => ⟨ c, conv v'', convE e, h, k, s[r:= HAN (comp y r c) (convE e) h] ⟩ 358 | | None => ⟪ Some r, MARKED :: k, s[r:= HAN (comp y r c) (convE e) h] ⟫ 359 | end. 360 | <== {eapply vm_unmark} 361 | match m with 362 | | Some v'' => ⟨ UNMARK c, conv v'', convE e, Some r, MARKED :: k, s[r:= HAN (comp y r c) (convE e) h] ⟩ 363 | | None => ⟪ Some r, MARKED :: k, s[r:= HAN (comp y r c) (convE e) h] ⟫ 364 | end. 365 | <|= {apply IHE1} 366 | ⟨ comp x (next r) (UNMARK c), a, convE e, Some r, MARKED :: k, s[r:= HAN (comp y r c) (convE e) h] ⟩. 367 | <== {apply vm_mark} 368 | ⟨ MARK r (comp y r c) (comp x (next r) (UNMARK c)), a, convE e, h, k, s ⟩. 369 | []. 370 | 371 | 372 | 373 | (** - [Var i ⇓[e] v] *) 374 | 375 | begin 376 | match nth e i with 377 | | Some v' => ⟨c, conv v', convE e, h, k, s⟩ 378 | | None => ⟪ h, k, s ⟫ 379 | end. 380 | <== {apply vm_lookup_fail; unfold convE; rewrite nth_map; rewr_assumption} 381 | match nth e i with 382 | | Some v' => ⟨c, conv v', convE e, h, k, s⟩ 383 | | None => ⟨LOOKUP i c, a , convE e, h, k, s ⟩ 384 | end. 385 | <== {apply vm_lookup; unfold convE; rewrite nth_map; rewr_assumption} 386 | ⟨LOOKUP i c, a , convE e, h, k, s ⟩. 387 | []. 388 | 389 | (** - [Abs x ⇓[e] Clo x e] *) 390 | 391 | begin 392 | ⟨c, Clo' (comp x (next first) RET) (convE e), convE e, h, k, s ⟩. 393 | <== { apply vm_abs } 394 | ⟨ABS (comp x (next first) RET) c, a, convE e, h, k, s ⟩. 395 | []. 396 | 397 | (** - [App x y ⇓[e] x''] *) 398 | 399 | begin 400 | match x'' with 401 | | Some v' => ⟨ c, conv v', convE e, h, k, s ⟩ 402 | | None => ⟪ h, k, s ⟫ 403 | end. 404 | <== { apply vm_ret } 405 | match x'' with 406 | | Some v' => ⟨RET, conv v', convE (y' :: e'), h, MEM s :: k, empty[first:=CLO c (convE e)]⟩ 407 | | None => ⟪ h, k, s ⟫ 408 | end. 409 | <== { apply vm_fail_pop } 410 | match x'' with 411 | | Some v' => ⟨RET, conv v', convE (y' :: e'), h, MEM s :: k, empty[first:=CLO c (convE e)]⟩ 412 | | None => ⟪ h, MEM s :: k, empty[first:=CLO c (convE e)] ⟫ 413 | end. 414 | <|= {apply IHE3} 415 | ⟨comp x' (next first) RET, Num' 0, convE (y' :: e'), h, MEM s :: k, empty[first:=CLO c (convE e)]⟩. 416 | = {auto} 417 | ⟨comp x' (next first) RET, Num' 0, conv y' :: convE e', h, MEM s:: k, empty[first:=CLO c (convE e)]⟩. 418 | <== {apply_eq vm_app;try rewrite truncate_set} 419 | ⟨APP r c, conv y', convE e, h, k, s[r:=CLO (comp x' (next first) RET) (convE e')]⟩. 420 | <|= {apply IHE2} 421 | ⟨comp y (next r) (APP r c), (Clo' (comp x' (next first) RET) (convE e')), convE e, h, k, s[r:=CLO (comp x' (next first) RET) (convE e')]⟩. 422 | <== { apply vm_stc } 423 | ⟨STC r (comp y (next r) (APP r c)), (Clo' (comp x' (next first) RET) (convE e')), convE e, h, k, s⟩. 424 | = {auto} 425 | ⟨STC r (comp y (next r) (APP r c)), conv (Clo x' e'), convE e, h, k, s ⟩. 426 | <|= { apply IHE1 } 427 | ⟨comp x r (STC r (comp y (next r) (APP r c))), a, convE e, h, k, s ⟩. 428 | []. 429 | 430 | begin 431 | ⟪ h, k, s ⟫. 432 | = {auto} 433 | match x' with 434 | | Some (Clo x'' e') => match y' with 435 | | Some v => ⟨ APP r c, conv v, convE e, h, k, s[ r := CLO (comp x'' (next first) RET) (map conv e')] ⟩ 436 | | None => ⟪h, k, s ⟫ 437 | end 438 | | _ => ⟪h, k, s ⟫ 439 | end. 440 | ⊑ {auto} 441 | match x' with 442 | | Some (Clo x'' e') => match y' with 443 | | Some v => ⟨ APP r c, conv v, convE e, h, k, s[ r := CLO (comp x'' (next first) RET) (map conv e')] ⟩ 444 | | None => ⟪h, k, s[ r := CLO (comp x'' (next first) RET) (map conv e')] ⟫ 445 | end 446 | | _ => ⟪h, k, s ⟫ 447 | end. 448 | <|= {apply IHE2} 449 | match x' with 450 | | Some (Clo x'' e') => ⟨comp y (next r) (APP r c), conv (Clo x'' e'), convE e, h, k, s[ r := CLO (comp x'' (next first) RET) (map conv e')] ⟩ 451 | | _ => ⟪h, k, s ⟫ 452 | end. 453 | = {reflexivity} 454 | match x' with 455 | | Some v => match v with 456 | | Clo x'' e' => ⟨comp y (next r) (APP r c), conv (Clo x'' e'), convE e, h, k, 457 | s[ r := CLO (comp x'' (next first) RET) (map conv e')] ⟩ 458 | | _ => ⟪h, k, s ⟫ 459 | end 460 | | _ => ⟪h, k, s ⟫ 461 | end. 462 | <== {apply vm_stc} 463 | match x' with 464 | | Some v => match v with 465 | | Clo x'' e' => ⟨STC r (comp y (next r) (APP r c)), conv (Clo x'' e'), convE e, h, k, s ⟩ 466 | | _ => ⟪h, k, s ⟫ 467 | end 468 | | _ => ⟪h, k, s ⟫ 469 | end. 470 | <== {apply vm_stc_fail} 471 | match x' with 472 | | Some v => ⟨STC r (comp y (next r) (APP r c)), conv v, convE e, h, k, s ⟩ 473 | | _ => ⟪h, k, s ⟫ 474 | end. 475 | <|= {apply IHE1} 476 | ⟨comp x r (STC r (comp y (next r) (APP r c))), a, convE e, h, k, s ⟩. 477 | []. 478 | 479 | Qed. 480 | 481 | (** * Soundness *) 482 | 483 | Lemma determ_vm : determ VM. 484 | intros C C1 C2 V. induction V; intro V'; inversion V'; subst; congruence. 485 | Qed. 486 | 487 | 488 | Definition terminates (p : Expr) : Prop := exists r, p ⇓[nil] Some r. 489 | 490 | Theorem sound p a s h C : freeFrom first s -> terminates p -> ⟨compile p, a, nil, h, nil, s⟩ =>>! C -> 491 | exists v s', C = ⟨HALT , conv v, nil, h, nil, s'⟩ /\ p ⇓[nil] Some v. 492 | Proof. 493 | unfold terminates. intros F T M. destruct T as [v T]. 494 | pose (spec p (Some v) nil first HALT a s nil h F T) as H'. 495 | unfold Reach in *. repeat autodestruct. 496 | pose (determ_trc determ_vm) as D. 497 | unfold determ in D. inversion H0. inversion H5. subst. 498 | exists v. eexists. split. eapply D. apply M. split. 499 | unfold compile. 500 | simpl in *. apply H. intro Contra. destruct Contra. 501 | inversion H1. assumption. 502 | Qed. 503 | 504 | Example test1 := (Catch (App (Abs (Throw)) (Val 2)) (Val 3)). 505 | 506 | Example test1_eval : exists v, test1 ⇓[ nil ] v. 507 | unfold test1. 508 | eexists. 509 | apply eval_catch. 510 | eapply eval_app. 511 | eapply eval_abs. 512 | eapply eval_val. 513 | eapply eval_throw. 514 | eapply eval_val. 515 | Qed. 516 | 517 | (* Compute (compile test1). *) 518 | 519 | Example test1_vm h : exists s a, ⟨(compile (Catch (App (Abs (Throw)) (Val 2)) (Val 3))), a, nil, h, nil, empty⟩ 520 | =>> ⟨HALT, a, nil, h, nil, s⟩. 521 | Proof. 522 | eexists. eexists. 523 | unfold compile. simpl. 524 | eapply trc_step_trans'. 525 | eapply vm_mark. 526 | eapply trc_step_trans'. 527 | eapply vm_abs. 528 | eapply trc_step_trans'. 529 | eapply vm_stc. 530 | eapply trc_step_trans'. 531 | eapply vm_load. 532 | eapply trc_step_trans'. 533 | eapply vm_app. 534 | rewrite get_set. reflexivity. 535 | eapply trc_step_trans'. 536 | eapply vm_throw. 537 | eapply trc_step_trans'. 538 | eapply vm_fail_pop. 539 | eapply trc_step_trans'. 540 | eapply vm_fail. rewrite truncate_set. rewrite get_set. reflexivity. apply freeFrom_set. auto with memory. 541 | eapply trc_step_trans'. 542 | eapply vm_load. 543 | eapply trc_refl. 544 | Qed. 545 | 546 | 547 | End Lambda. 548 | -------------------------------------------------------------------------------- /coq/Makefile: -------------------------------------------------------------------------------- 1 | ############################################################################### 2 | ## v # The Coq Proof Assistant ## 3 | ## /dev/null 2>/dev/null; echo $$?)) 72 | STDTIME?=command time -f $(TIMEFMT) 73 | else 74 | ifeq (0,$(shell gtime -f $(TIMEFMT) true >/dev/null 2>/dev/null; echo $$?)) 75 | STDTIME?=gtime -f $(TIMEFMT) 76 | else 77 | STDTIME?=command time 78 | endif 79 | endif 80 | else 81 | STDTIME?=command time -f $(TIMEFMT) 82 | endif 83 | 84 | # Coq binaries 85 | COQC ?= "$(COQBIN)coqc" 86 | COQTOP ?= "$(COQBIN)coqtop" 87 | COQCHK ?= "$(COQBIN)coqchk" 88 | COQDEP ?= "$(COQBIN)coqdep" 89 | COQDOC ?= "$(COQBIN)coqdoc" 90 | COQMKFILE ?= "$(COQBIN)coq_makefile" 91 | 92 | # Timing scripts 93 | COQMAKE_ONE_TIME_FILE ?= "$(COQLIB)/tools/make-one-time-file.py" 94 | COQMAKE_BOTH_TIME_FILES ?= "$(COQLIB)/tools/make-both-time-files.py" 95 | COQMAKE_BOTH_SINGLE_TIMING_FILES ?= "$(COQLIB)/tools/make-both-single-timing-files.py" 96 | BEFORE ?= 97 | AFTER ?= 98 | 99 | # FIXME this should be generated by Coq (modules already linked by Coq) 100 | CAMLDONTLINK=camlp5.gramlib,unix,str 101 | 102 | # OCaml binaries 103 | CAMLC ?= "$(OCAMLFIND)" ocamlc -c 104 | CAMLOPTC ?= "$(OCAMLFIND)" opt -c 105 | CAMLLINK ?= "$(OCAMLFIND)" ocamlc -linkpkg -dontlink $(CAMLDONTLINK) 106 | CAMLOPTLINK ?= "$(OCAMLFIND)" opt -linkpkg -dontlink $(CAMLDONTLINK) 107 | CAMLDOC ?= "$(OCAMLFIND)" ocamldoc 108 | CAMLDEP ?= "$(OCAMLFIND)" ocamldep -slash -ml-synonym .ml4 -ml-synonym .mlpack 109 | 110 | # DESTDIR is prepended to all installation paths 111 | DESTDIR ?= 112 | 113 | # Debug builds, typically -g to OCaml, -debug to Coq. 114 | CAMLDEBUG ?= 115 | COQDEBUG ?= 116 | 117 | # Extra packages to be linked in (as in findlib -package) 118 | CAMLPKGS ?= 119 | 120 | # Option for making timing files 121 | TIMING?= 122 | # Option for changing sorting of timing output file 123 | TIMING_SORT_BY ?= auto 124 | # Output file names for timed builds 125 | TIME_OF_BUILD_FILE ?= time-of-build.log 126 | TIME_OF_BUILD_BEFORE_FILE ?= time-of-build-before.log 127 | TIME_OF_BUILD_AFTER_FILE ?= time-of-build-after.log 128 | TIME_OF_PRETTY_BUILD_FILE ?= time-of-build-pretty.log 129 | TIME_OF_PRETTY_BOTH_BUILD_FILE ?= time-of-build-both.log 130 | TIME_OF_PRETTY_BUILD_EXTRA_FILES ?= - # also output to the command line 131 | 132 | ########## End of parameters ################################################## 133 | # What follows may be relevant to you only if you need to 134 | # extend this Makefile. If so, look for 'Extension point' here and 135 | # put in makefile.local double colon rules accordingly. 136 | # E.g. to perform some work after the all target completes you can write 137 | # 138 | # post-all:: 139 | # echo "All done!" 140 | # 141 | # in makefile.local 142 | # 143 | ############################################################################### 144 | 145 | 146 | 147 | 148 | # Flags ####################################################################### 149 | # 150 | # We define a bunch of variables combining the parameters. 151 | # To add additional flags to coq, coqchk or coqdoc, set the 152 | # {COQ,COQCHK,COQDOC}EXTRAFLAGS variable to whatever you want to add. 153 | # To overwrite the default choice and set your own flags entirely, set the 154 | # {COQ,COQCHK,COQDOC}FLAGS variable. 155 | 156 | SHOW := $(if $(VERBOSE),@true "",@echo "") 157 | HIDE := $(if $(VERBOSE),,@) 158 | 159 | TIMER=$(if $(TIMED), $(STDTIME), $(TIMECMD)) 160 | 161 | OPT?= 162 | 163 | # The DYNOBJ and DYNLIB variables are used by "coqdep -dyndep var" in .v.d 164 | ifeq '$(OPT)' '-byte' 165 | USEBYTE:=true 166 | DYNOBJ:=.cma 167 | DYNLIB:=.cma 168 | else 169 | USEBYTE:= 170 | DYNOBJ:=.cmxs 171 | DYNLIB:=.cmxs 172 | endif 173 | 174 | # these variables are meant to be overriden if you want to add *extra* flags 175 | COQEXTRAFLAGS?= 176 | COQCHKEXTRAFLAGS?= 177 | COQDOCEXTRAFLAGS?= 178 | 179 | # these flags do NOT contain the libraries, to make them easier to overwrite 180 | COQFLAGS?=-q $(OPT) $(OTHERFLAGS) $(COQEXTRAFLAGS) 181 | COQCHKFLAGS?=-silent -o $(COQCHKEXTRAFLAGS) 182 | COQDOCFLAGS?=-interpolate -utf8 $(COQDOCEXTRAFLAGS) 183 | 184 | COQDOCLIBS?=$(COQLIBS_NOML) 185 | 186 | # The version of Coq being run and the version of coq_makefile that 187 | # generated this makefile 188 | COQ_VERSION:=$(shell $(COQC) --print-version | cut -d " " -f 1) 189 | COQMAKEFILE_VERSION:=8.9.0 190 | 191 | COQSRCLIBS?= $(foreach d,$(COQ_SRC_SUBDIRS), -I "$(COQLIB)$(d)") 192 | 193 | CAMLFLAGS+=$(OCAMLLIBS) $(COQSRCLIBS) -I $(CAMLP5LIB) 194 | 195 | # ocamldoc fails with unknown argument otherwise 196 | CAMLDOCFLAGS=$(filter-out -annot, $(filter-out -bin-annot, $(CAMLFLAGS))) 197 | 198 | # FIXME This should be generated by Coq 199 | GRAMMARS:=grammar.cma 200 | CAMLP5EXTEND=pa_extend.cmo q_MLast.cmo pa_macro.cmo 201 | 202 | CAMLLIB:=$(shell "$(OCAMLFIND)" printconf stdlib 2> /dev/null) 203 | ifeq (,$(CAMLLIB)) 204 | PP=$(error "Cannot find the 'ocamlfind' binary used to build Coq ($(OCAMLFIND)). Pre-compiled binary packages of Coq do not support compiling plugins this way. Please download the sources of Coq and run the Windows build script.") 205 | else 206 | PP:=-pp '$(CAMLP5O) -I $(CAMLLIB) -I "$(COQLIB)/grammar" $(CAMLP5EXTEND) $(GRAMMARS) $(CAMLP5OPTIONS) -impl' 207 | endif 208 | 209 | ifneq (,$(TIMING)) 210 | TIMING_ARG=-time 211 | ifeq (after,$(TIMING)) 212 | TIMING_EXT=after-timing 213 | else 214 | ifeq (before,$(TIMING)) 215 | TIMING_EXT=before-timing 216 | else 217 | TIMING_EXT=timing 218 | endif 219 | endif 220 | else 221 | TIMING_ARG= 222 | endif 223 | 224 | # Retro compatibility (DESTDIR is standard on Unix, DSTROOT is not) 225 | ifdef DSTROOT 226 | DESTDIR := $(DSTROOT) 227 | endif 228 | 229 | concat_path = $(if $(1),$(1)/$(if $(COQMF_WINDRIVE),$(subst $(COQMF_WINDRIVE),/,$(2)),$(2)),$(2)) 230 | 231 | COQLIBINSTALL = $(call concat_path,$(DESTDIR),$(COQLIB)user-contrib) 232 | COQDOCINSTALL = $(call concat_path,$(DESTDIR),$(DOCDIR)user-contrib) 233 | COQTOPINSTALL = $(call concat_path,$(DESTDIR),$(COQLIB)toploop) 234 | 235 | # Files ####################################################################### 236 | # 237 | # We here define a bunch of variables about the files being part of the 238 | # Coq project in order to ease the writing of build target and build rules 239 | 240 | VDFILE := .coqdeps 241 | 242 | ALLSRCFILES := \ 243 | $(ML4FILES) \ 244 | $(MLFILES) \ 245 | $(MLPACKFILES) \ 246 | $(MLLIBFILES) \ 247 | $(MLIFILES) 248 | 249 | # helpers 250 | vo_to_obj = $(addsuffix .o,\ 251 | $(filter-out Warning: Error:,\ 252 | $(shell $(COQTOP) -q -noinit -batch -quiet -print-mod-uid $(1)))) 253 | strip_dotslash = $(patsubst ./%,%,$(1)) 254 | VO = vo 255 | 256 | VOFILES = $(VFILES:.v=.$(VO)) 257 | GLOBFILES = $(VFILES:.v=.glob) 258 | HTMLFILES = $(VFILES:.v=.html) 259 | GHTMLFILES = $(VFILES:.v=.g.html) 260 | BEAUTYFILES = $(addsuffix .beautified,$(VFILES)) 261 | TEXFILES = $(VFILES:.v=.tex) 262 | GTEXFILES = $(VFILES:.v=.g.tex) 263 | CMOFILES = \ 264 | $(ML4FILES:.ml4=.cmo) \ 265 | $(MLFILES:.ml=.cmo) \ 266 | $(MLPACKFILES:.mlpack=.cmo) 267 | CMXFILES = $(CMOFILES:.cmo=.cmx) 268 | OFILES = $(CMXFILES:.cmx=.o) 269 | CMAFILES = $(MLLIBFILES:.mllib=.cma) $(MLPACKFILES:.mlpack=.cma) 270 | CMXAFILES = $(CMAFILES:.cma=.cmxa) 271 | CMIFILES = \ 272 | $(CMOFILES:.cmo=.cmi) \ 273 | $(MLIFILES:.mli=.cmi) 274 | # the /if/ is because old _CoqProject did not list a .ml(pack|lib) but just 275 | # a .ml4 file 276 | CMXSFILES = \ 277 | $(MLPACKFILES:.mlpack=.cmxs) \ 278 | $(CMXAFILES:.cmxa=.cmxs) \ 279 | $(if $(MLPACKFILES)$(CMXAFILES),,\ 280 | $(ML4FILES:.ml4=.cmxs) $(MLFILES:.ml=.cmxs)) 281 | 282 | # files that are packed into a plugin (no extension) 283 | PACKEDFILES = \ 284 | $(call strip_dotslash, \ 285 | $(foreach lib, \ 286 | $(call strip_dotslash, \ 287 | $(MLPACKFILES:.mlpack=_MLPACK_DEPENDENCIES)),$($(lib)))) 288 | # files that are archived into a .cma (mllib) 289 | LIBEDFILES = \ 290 | $(call strip_dotslash, \ 291 | $(foreach lib, \ 292 | $(call strip_dotslash, \ 293 | $(MLLIBFILES:.mllib=_MLLIB_DEPENDENCIES)),$($(lib)))) 294 | CMIFILESTOINSTALL = $(filter-out $(addsuffix .cmi,$(PACKEDFILES)),$(CMIFILES)) 295 | CMOFILESTOINSTALL = $(filter-out $(addsuffix .cmo,$(PACKEDFILES)),$(CMOFILES)) 296 | OBJFILES = $(call vo_to_obj,$(VOFILES)) 297 | ALLNATIVEFILES = \ 298 | $(OBJFILES:.o=.cmi) \ 299 | $(OBJFILES:.o=.cmx) \ 300 | $(OBJFILES:.o=.cmxs) 301 | # trick: wildcard filters out non-existing files, so that `install` doesn't show 302 | # warnings and `clean` doesn't pass to rm a list of files that is too long for 303 | # the shell. 304 | NATIVEFILES = $(wildcard $(ALLNATIVEFILES)) 305 | FILESTOINSTALL = \ 306 | $(VOFILES) \ 307 | $(VFILES) \ 308 | $(GLOBFILES) \ 309 | $(NATIVEFILES) \ 310 | $(CMIFILESTOINSTALL) 311 | BYTEFILESTOINSTALL = \ 312 | $(CMOFILESTOINSTALL) \ 313 | $(CMAFILES) 314 | ifeq '$(HASNATDYNLINK)' 'true' 315 | DO_NATDYNLINK = yes 316 | FILESTOINSTALL += $(CMXSFILES) $(CMXAFILES) $(CMOFILESTOINSTALL:.cmo=.cmx) 317 | else 318 | DO_NATDYNLINK = 319 | endif 320 | 321 | ALLDFILES = $(addsuffix .d,$(ALLSRCFILES) $(VDFILE)) 322 | 323 | # Compilation targets ######################################################### 324 | 325 | all: 326 | $(HIDE)$(MAKE) --no-print-directory -f "$(SELF)" pre-all 327 | $(HIDE)$(MAKE) --no-print-directory -f "$(SELF)" real-all 328 | $(HIDE)$(MAKE) --no-print-directory -f "$(SELF)" post-all 329 | .PHONY: all 330 | 331 | all.timing.diff: 332 | $(HIDE)$(MAKE) --no-print-directory -f "$(SELF)" pre-all 333 | $(HIDE)$(MAKE) --no-print-directory -f "$(SELF)" real-all.timing.diff TIME_OF_PRETTY_BUILD_EXTRA_FILES="" 334 | $(HIDE)$(MAKE) --no-print-directory -f "$(SELF)" post-all 335 | .PHONY: all.timing.diff 336 | 337 | make-pretty-timed-before:: TIME_OF_BUILD_FILE=$(TIME_OF_BUILD_BEFORE_FILE) 338 | make-pretty-timed-after:: TIME_OF_BUILD_FILE=$(TIME_OF_BUILD_AFTER_FILE) 339 | make-pretty-timed make-pretty-timed-before make-pretty-timed-after:: 340 | $(HIDE)rm -f pretty-timed-success.ok 341 | $(HIDE)($(MAKE) --no-print-directory -f "$(PARENT)" $(TGTS) TIMED=1 2>&1 && touch pretty-timed-success.ok) | tee -a $(TIME_OF_BUILD_FILE) 342 | $(HIDE)rm pretty-timed-success.ok # must not be -f; must fail if the touch failed 343 | print-pretty-timed:: 344 | $(HIDE)$(COQMAKE_ONE_TIME_FILE) $(TIME_OF_BUILD_FILE) $(TIME_OF_PRETTY_BUILD_FILE) $(TIME_OF_PRETTY_BUILD_EXTRA_FILES) 345 | print-pretty-timed-diff:: 346 | $(HIDE)$(COQMAKE_BOTH_TIME_FILES) --sort-by=$(TIMING_SORT_BY) $(TIME_OF_BUILD_AFTER_FILE) $(TIME_OF_BUILD_BEFORE_FILE) $(TIME_OF_PRETTY_BOTH_BUILD_FILE) $(TIME_OF_PRETTY_BUILD_EXTRA_FILES) 347 | ifeq (,$(BEFORE)) 348 | print-pretty-single-time-diff:: 349 | @echo 'Error: Usage: $(MAKE) print-pretty-single-time-diff AFTER=path/to/file.v.after-timing BEFORE=path/to/file.v.before-timing' 350 | $(HIDE)false 351 | else 352 | ifeq (,$(AFTER)) 353 | print-pretty-single-time-diff:: 354 | @echo 'Error: Usage: $(MAKE) print-pretty-single-time-diff AFTER=path/to/file.v.after-timing BEFORE=path/to/file.v.before-timing' 355 | $(HIDE)false 356 | else 357 | print-pretty-single-time-diff:: 358 | $(HIDE)$(COQMAKE_BOTH_SINGLE_TIMING_FILES) --sort-by=$(TIMING_SORT_BY) $(AFTER) $(BEFORE) $(TIME_OF_PRETTY_BUILD_FILE) $(TIME_OF_PRETTY_BUILD_EXTRA_FILES) 359 | endif 360 | endif 361 | pretty-timed: 362 | $(HIDE)$(MAKE) --no-print-directory -f "$(PARENT)" make-pretty-timed 363 | $(HIDE)$(MAKE) --no-print-directory -f "$(SELF)" print-pretty-timed 364 | .PHONY: pretty-timed make-pretty-timed make-pretty-timed-before make-pretty-timed-after print-pretty-timed print-pretty-timed-diff print-pretty-single-time-diff 365 | 366 | # Extension points for actions to be performed before/after the all target 367 | pre-all:: 368 | @# Extension point 369 | $(HIDE)if [ "$(COQMAKEFILE_VERSION)" != "$(COQ_VERSION)" ]; then\ 370 | echo "W: This Makefile was generated by Coq $(COQMAKEFILE_VERSION)";\ 371 | echo "W: while the current Coq version is $(COQ_VERSION)";\ 372 | fi 373 | .PHONY: pre-all 374 | 375 | post-all:: 376 | @# Extension point 377 | .PHONY: post-all 378 | 379 | real-all: $(VOFILES) $(if $(USEBYTE),bytefiles,optfiles) 380 | .PHONY: real-all 381 | 382 | real-all.timing.diff: $(VOFILES:.vo=.v.timing.diff) 383 | .PHONY: real-all.timing.diff 384 | 385 | bytefiles: $(CMOFILES) $(CMAFILES) 386 | .PHONY: bytefiles 387 | 388 | optfiles: $(if $(DO_NATDYNLINK),$(CMXSFILES)) 389 | .PHONY: optfiles 390 | 391 | # FIXME, see Ralf's bugreport 392 | quick: $(VOFILES:.vo=.vio) 393 | .PHONY: quick 394 | 395 | vio2vo: 396 | $(TIMER) $(COQC) $(COQDEBUG) $(COQFLAGS) $(COQLIBS) \ 397 | -schedule-vio2vo $(J) $(VOFILES:%.vo=%.vio) 398 | .PHONY: vio2vo 399 | 400 | quick2vo: 401 | $(HIDE)make -j $(J) quick 402 | $(HIDE)VIOFILES=$$(for vofile in $(VOFILES); do \ 403 | viofile="$$(echo "$$vofile" | sed "s/\.vo$$/.vio/")"; \ 404 | if [ "$$vofile" -ot "$$viofile" -o ! -e "$$vofile" ]; then printf "$$viofile "; fi; \ 405 | done); \ 406 | echo "VIO2VO: $$VIOFILES"; \ 407 | if [ -n "$$VIOFILES" ]; then \ 408 | $(TIMER) $(COQC) $(COQDEBUG) $(COQFLAGS) $(COQLIBS) -schedule-vio2vo $(J) $$VIOFILES; \ 409 | fi 410 | .PHONY: quick2vo 411 | 412 | checkproofs: 413 | $(TIMER) $(COQC) $(COQDEBUG) $(COQFLAGS) $(COQLIBS) \ 414 | -schedule-vio-checking $(J) $(VOFILES:%.vo=%.vio) 415 | .PHONY: checkproofs 416 | 417 | validate: $(VOFILES) 418 | $(TIMER) $(COQCHK) $(COQCHKFLAGS) $(COQLIBS) $^ 419 | .PHONY: validate 420 | 421 | only: $(TGTS) 422 | .PHONY: only 423 | 424 | # Documentation targets ####################################################### 425 | 426 | html: $(GLOBFILES) $(VFILES) 427 | $(SHOW)'COQDOC -d html $(GAL)' 428 | $(HIDE)mkdir -p html 429 | $(HIDE)$(COQDOC) \ 430 | -toc $(COQDOCFLAGS) -html $(GAL) $(COQDOCLIBS) -d html $(VFILES) 431 | 432 | mlihtml: $(MLIFILES:.mli=.cmi) 433 | $(SHOW)'CAMLDOC -d $@' 434 | $(HIDE)mkdir $@ || rm -rf $@/* 435 | $(HIDE)$(CAMLDOC) -html \ 436 | -d $@ -m A $(CAMLDEBUG) $(CAMLDOCFLAGS) $(MLIFILES) 437 | 438 | all-mli.tex: $(MLIFILES:.mli=.cmi) 439 | $(SHOW)'CAMLDOC -latex $@' 440 | $(HIDE)$(CAMLDOC) -latex \ 441 | -o $@ -m A $(CAMLDEBUG) $(CAMLDOCFLAGS) $(MLIFILES) 442 | 443 | all.ps: $(VFILES) 444 | $(SHOW)'COQDOC -ps $(GAL)' 445 | $(HIDE)$(COQDOC) \ 446 | -toc $(COQDOCFLAGS) -ps $(GAL) $(COQDOCLIBS) \ 447 | -o $@ `$(COQDEP) -sort -suffix .v $(VFILES)` 448 | 449 | all.pdf: $(VFILES) 450 | $(SHOW)'COQDOC -pdf $(GAL)' 451 | $(HIDE)$(COQDOC) \ 452 | -toc $(COQDOCFLAGS) -pdf $(GAL) $(COQDOCLIBS) \ 453 | -o $@ `$(COQDEP) -sort -suffix .v $(VFILES)` 454 | 455 | # FIXME: not quite right, since the output name is different 456 | gallinahtml: GAL=-g 457 | gallinahtml: html 458 | 459 | all-gal.ps: GAL=-g 460 | all-gal.ps: all.ps 461 | 462 | all-gal.pdf: GAL=-g 463 | all-gal.pdf: all.pdf 464 | 465 | # ? 466 | beautify: $(BEAUTYFILES) 467 | for file in $^; do mv $${file%.beautified} $${file%beautified}old && mv $${file} $${file%.beautified}; done 468 | @echo 'Do not do "make clean" until you are sure that everything went well!' 469 | @echo 'If there were a problem, execute "for file in $$(find . -name \*.v.old -print); do mv $${file} $${file%.old}; done" in your shell/' 470 | .PHONY: beautify 471 | 472 | # Installation targets ######################################################## 473 | # 474 | # There rules can be extended in makefile.local 475 | # Extensions can't assume when they run. 476 | 477 | install: 478 | $(HIDE)for f in $(FILESTOINSTALL); do\ 479 | df="`$(COQMKFILE) -destination-of "$$f" $(COQLIBS)`";\ 480 | if [ "$$?" != "0" -o -z "$$df" ]; then\ 481 | echo SKIP "$$f" since it has no logical path;\ 482 | else\ 483 | install -d "$(COQLIBINSTALL)/$$df" &&\ 484 | install -m 0644 "$$f" "$(COQLIBINSTALL)/$$df" &&\ 485 | echo INSTALL "$$f" "$(COQLIBINSTALL)/$$df";\ 486 | fi;\ 487 | done 488 | $(HIDE)$(MAKE) install-extra -f "$(SELF)" 489 | install-extra:: 490 | @# Extension point 491 | .PHONY: install install-extra 492 | 493 | install-byte: 494 | $(HIDE)for f in $(BYTEFILESTOINSTALL); do\ 495 | df="`$(COQMKFILE) -destination-of "$$f" $(COQLIBS)`";\ 496 | if [ "$$?" != "0" -o -z "$$df" ]; then\ 497 | echo SKIP "$$f" since it has no logical path;\ 498 | else\ 499 | install -d "$(COQLIBINSTALL)/$$df" &&\ 500 | install -m 0644 "$$f" "$(COQLIBINSTALL)/$$df" &&\ 501 | echo INSTALL "$$f" "$(COQLIBINSTALL)/$$df";\ 502 | fi;\ 503 | done 504 | 505 | install-doc:: html mlihtml 506 | @# Extension point 507 | $(HIDE)install -d "$(COQDOCINSTALL)/$(INSTALLCOQDOCROOT)/html" 508 | $(HIDE)for i in html/*; do \ 509 | dest="$(COQDOCINSTALL)/$(INSTALLCOQDOCROOT)/$$i";\ 510 | install -m 0644 "$$i" "$$dest";\ 511 | echo INSTALL "$$i" "$$dest";\ 512 | done 513 | $(HIDE)install -d \ 514 | "$(COQDOCINSTALL)/$(INSTALLCOQDOCROOT)/mlihtml" 515 | $(HIDE)for i in mlihtml/*; do \ 516 | dest="$(COQDOCINSTALL)/$(INSTALLCOQDOCROOT)/$$i";\ 517 | install -m 0644 "$$i" "$$dest";\ 518 | echo INSTALL "$$i" "$$dest";\ 519 | done 520 | .PHONY: install-doc 521 | 522 | uninstall:: 523 | @# Extension point 524 | $(HIDE)for f in $(FILESTOINSTALL); do \ 525 | df="`$(COQMKFILE) -destination-of "$$f" $(COQLIBS)`" &&\ 526 | instf="$(COQLIBINSTALL)/$$df/`basename $$f`" &&\ 527 | rm -f "$$instf" &&\ 528 | echo RM "$$instf" &&\ 529 | (rmdir "$(call concat_path,,$(COQLIBINSTALL)/$$df/)" 2>/dev/null || true); \ 530 | done 531 | .PHONY: uninstall 532 | 533 | uninstall-doc:: 534 | @# Extension point 535 | $(SHOW)'RM $(COQDOCINSTALL)/$(INSTALLCOQDOCROOT)/html' 536 | $(HIDE)rm -rf "$(COQDOCINSTALL)/$(INSTALLCOQDOCROOT)/html" 537 | $(SHOW)'RM $(COQDOCINSTALL)/$(INSTALLCOQDOCROOT)/mlihtml' 538 | $(HIDE)rm -rf "$(COQDOCINSTALL)/$(INSTALLCOQDOCROOT)/mlihtml" 539 | $(HIDE) rmdir "$(COQDOCINSTALL)/$(INSTALLCOQDOCROOT)/" || true 540 | .PHONY: uninstall-doc 541 | 542 | # Cleaning #################################################################### 543 | # 544 | # There rules can be extended in makefile.local 545 | # Extensions can't assume when they run. 546 | 547 | clean:: 548 | @# Extension point 549 | $(SHOW)'CLEAN' 550 | $(HIDE)rm -f $(CMOFILES) 551 | $(HIDE)rm -f $(CMIFILES) 552 | $(HIDE)rm -f $(CMAFILES) 553 | $(HIDE)rm -f $(CMOFILES:.cmo=.cmx) 554 | $(HIDE)rm -f $(CMXAFILES) 555 | $(HIDE)rm -f $(CMXSFILES) 556 | $(HIDE)rm -f $(CMOFILES:.cmo=.o) 557 | $(HIDE)rm -f $(CMXAFILES:.cmxa=.a) 558 | $(HIDE)rm -f $(ALLDFILES) 559 | $(HIDE)rm -f $(NATIVEFILES) 560 | $(HIDE)find . -name .coq-native -type d -empty -delete 561 | $(HIDE)rm -f $(VOFILES) 562 | $(HIDE)rm -f $(VOFILES:.vo=.vio) 563 | $(HIDE)rm -f $(BEAUTYFILES) $(VFILES:=.old) 564 | $(HIDE)rm -f all.ps all-gal.ps all.pdf all-gal.pdf all.glob all-mli.tex 565 | $(HIDE)rm -f $(VFILES:.v=.glob) 566 | $(HIDE)rm -f $(VFILES:.v=.tex) 567 | $(HIDE)rm -f $(VFILES:.v=.g.tex) 568 | $(HIDE)rm -f pretty-timed-success.ok 569 | $(HIDE)rm -rf html mlihtml 570 | .PHONY: clean 571 | 572 | cleanall:: clean 573 | @# Extension point 574 | $(SHOW)'CLEAN *.aux *.timing' 575 | $(HIDE)rm -f $(foreach f,$(VFILES:.v=),$(dir $(f)).$(notdir $(f)).aux) 576 | $(HIDE)rm -f $(TIME_OF_BUILD_FILE) $(TIME_OF_BUILD_BEFORE_FILE) $(TIME_OF_BUILD_AFTER_FILE) $(TIME_OF_PRETTY_BUILD_FILE) $(TIME_OF_PRETTY_BOTH_BUILD_FILE) 577 | $(HIDE)rm -f $(VOFILES:.vo=.v.timing) 578 | $(HIDE)rm -f $(VOFILES:.vo=.v.before-timing) 579 | $(HIDE)rm -f $(VOFILES:.vo=.v.after-timing) 580 | $(HIDE)rm -f $(VOFILES:.vo=.v.timing.diff) 581 | .PHONY: cleanall 582 | 583 | archclean:: 584 | @# Extension point 585 | $(SHOW)'CLEAN *.cmx *.o' 586 | $(HIDE)rm -f $(NATIVEFILES) 587 | $(HIDE)rm -f $(CMOFILES:%.cmo=%.cmx) 588 | .PHONY: archclean 589 | 590 | 591 | # Compilation rules ########################################################### 592 | 593 | $(MLIFILES:.mli=.cmi): %.cmi: %.mli 594 | $(SHOW)'CAMLC -c $<' 595 | $(HIDE)$(CAMLC) $(CAMLDEBUG) $(CAMLFLAGS) $(CAMLPKGS) $< 596 | 597 | $(ML4FILES:.ml4=.cmo): %.cmo: %.ml4 598 | $(SHOW)'CAMLC -pp -c $<' 599 | $(HIDE)$(CAMLC) $(CAMLDEBUG) $(CAMLFLAGS) $(CAMLPKGS) $(PP) -impl $< 600 | 601 | $(ML4FILES:.ml4=.cmx): %.cmx: %.ml4 602 | $(SHOW)'CAMLOPT -pp -c $(FOR_PACK) $<' 603 | $(HIDE)$(CAMLOPTC) $(CAMLDEBUG) $(CAMLFLAGS) $(CAMLPKGS) $(PP) $(FOR_PACK) -impl $< 604 | 605 | $(MLFILES:.ml=.cmo): %.cmo: %.ml 606 | $(SHOW)'CAMLC -c $<' 607 | $(HIDE)$(CAMLC) $(CAMLDEBUG) $(CAMLFLAGS) $(CAMLPKGS) $< 608 | 609 | $(MLFILES:.ml=.cmx): %.cmx: %.ml 610 | $(SHOW)'CAMLOPT -c $(FOR_PACK) $<' 611 | $(HIDE)$(CAMLOPTC) $(CAMLDEBUG) $(CAMLFLAGS) $(CAMLPKGS) $(FOR_PACK) $< 612 | 613 | 614 | $(MLLIBFILES:.mllib=.cmxs): %.cmxs: %.cmxa 615 | $(SHOW)'CAMLOPT -shared -o $@' 616 | $(HIDE)$(CAMLOPTLINK) $(CAMLDEBUG) $(CAMLFLAGS) $(CAMLPKGS) \ 617 | -linkall -shared -o $@ $< 618 | 619 | $(MLLIBFILES:.mllib=.cma): %.cma: | %.mllib 620 | $(SHOW)'CAMLC -a -o $@' 621 | $(HIDE)$(CAMLLINK) $(CAMLDEBUG) $(CAMLFLAGS) $(CAMLPKGS) -a -o $@ $^ 622 | 623 | $(MLLIBFILES:.mllib=.cmxa): %.cmxa: | %.mllib 624 | $(SHOW)'CAMLOPT -a -o $@' 625 | $(HIDE)$(CAMLOPTLINK) $(CAMLDEBUG) $(CAMLFLAGS) $(CAMLPKGS) -a -o $@ $^ 626 | 627 | 628 | $(MLPACKFILES:.mlpack=.cmxs): %.cmxs: %.cmxa 629 | $(SHOW)'CAMLOPT -shared -o $@' 630 | $(HIDE)$(CAMLOPTLINK) $(CAMLDEBUG) $(CAMLFLAGS) $(CAMLPKGS) \ 631 | -shared -linkall -o $@ $< 632 | 633 | $(MLPACKFILES:.mlpack=.cmxa): %.cmxa: %.cmx 634 | $(SHOW)'CAMLOPT -a -o $@' 635 | $(HIDE)$(CAMLOPTLINK) $(CAMLDEBUG) $(CAMLFLAGS) -a -o $@ $< 636 | 637 | $(MLPACKFILES:.mlpack=.cma): %.cma: %.cmo | %.mlpack 638 | $(SHOW)'CAMLC -a -o $@' 639 | $(HIDE)$(CAMLLINK) $(CAMLDEBUG) $(CAMLFLAGS) $(CAMLPKGS) -a -o $@ $^ 640 | 641 | $(MLPACKFILES:.mlpack=.cmo): %.cmo: | %.mlpack 642 | $(SHOW)'CAMLC -pack -o $@' 643 | $(HIDE)$(CAMLLINK) $(CAMLDEBUG) $(CAMLFLAGS) -pack -o $@ $^ 644 | 645 | $(MLPACKFILES:.mlpack=.cmx): %.cmx: | %.mlpack 646 | $(SHOW)'CAMLOPT -pack -o $@' 647 | $(HIDE)$(CAMLOPTLINK) $(CAMLDEBUG) $(CAMLFLAGS) -pack -o $@ $^ 648 | 649 | # This rule is for _CoqProject with no .mllib nor .mlpack 650 | $(filter-out $(MLLIBFILES:.mllib=.cmxs) $(MLPACKFILES:.mlpack=.cmxs) $(addsuffix .cmxs,$(PACKEDFILES)) $(addsuffix .cmxs,$(LIBEDFILES)),$(MLFILES:.ml=.cmxs) $(ML4FILES:.ml4=.cmxs)): %.cmxs: %.cmx 651 | $(SHOW)'[deprecated,use-mllib-or-mlpack] CAMLOPT -shared -o $@' 652 | $(HIDE)$(CAMLOPTLINK) $(CAMLDEBUG) $(CAMLFLAGS) $(CAMLPKGS) \ 653 | -shared -o $@ $< 654 | 655 | ifneq (,$(TIMING)) 656 | TIMING_EXTRA = > $<.$(TIMING_EXT) 657 | else 658 | TIMING_EXTRA = 659 | endif 660 | 661 | $(VOFILES): %.vo: %.v 662 | $(SHOW)COQC $< 663 | $(HIDE)$(TIMER) $(COQC) $(COQDEBUG) $(TIMING_ARG) $(COQFLAGS) $(COQLIBS) $< $(TIMING_EXTRA) 664 | 665 | # FIXME ?merge with .vo / .vio ? 666 | $(GLOBFILES): %.glob: %.v 667 | $(TIMER) $(COQC) $(COQDEBUG) $(COQFLAGS) $(COQLIBS) $< 668 | 669 | $(VFILES:.v=.vio): %.vio: %.v 670 | $(SHOW)COQC -quick $< 671 | $(HIDE)$(TIMER) $(COQC) -quick $(COQDEBUG) $(COQFLAGS) $(COQLIBS) $< 672 | 673 | $(addsuffix .timing.diff,$(VFILES)): %.timing.diff : %.before-timing %.after-timing 674 | $(SHOW)PYTHON TIMING-DIFF $< 675 | $(HIDE)$(MAKE) --no-print-directory -f "$(SELF)" print-pretty-single-time-diff BEFORE=$*.before-timing AFTER=$*.after-timing TIME_OF_PRETTY_BUILD_FILE="$@" 676 | 677 | $(BEAUTYFILES): %.v.beautified: %.v 678 | $(SHOW)'BEAUTIFY $<' 679 | $(HIDE)$(TIMER) $(COQC) $(COQDEBUG) $(COQFLAGS) $(COQLIBS) -beautify $< 680 | 681 | $(TEXFILES): %.tex: %.v 682 | $(SHOW)'COQDOC -latex $<' 683 | $(HIDE)$(COQDOC) $(COQDOCFLAGS) -latex $< -o $@ 684 | 685 | $(GTEXFILES): %.g.tex: %.v 686 | $(SHOW)'COQDOC -latex -g $<' 687 | $(HIDE)$(COQDOC) $(COQDOCFLAGS) -latex -g $< -o $@ 688 | 689 | $(HTMLFILES): %.html: %.v %.glob 690 | $(SHOW)'COQDOC -html $<' 691 | $(HIDE)$(COQDOC) $(COQDOCFLAGS) -html $< -o $@ 692 | 693 | $(GHTMLFILES): %.g.html: %.v %.glob 694 | $(SHOW)'COQDOC -html -g $<' 695 | $(HIDE)$(COQDOC) $(COQDOCFLAGS) -html -g $< -o $@ 696 | 697 | # Dependency files ############################################################ 698 | 699 | ifneq ($(filter-out archclean clean cleanall printenv make-pretty-timed make-pretty-timed-before make-pretty-timed-after print-pretty-timed print-pretty-timed-diff print-pretty-single-time-diff,$(MAKECMDGOALS)),) 700 | -include $(ALLDFILES) 701 | else 702 | ifeq ($(MAKECMDGOALS),) 703 | -include $(ALLDFILES) 704 | endif 705 | endif 706 | 707 | .SECONDARY: $(ALLDFILES) 708 | 709 | redir_if_ok = > "$@" || ( RV=$$?; rm -f "$@"; exit $$RV ) 710 | 711 | $(addsuffix .d,$(MLIFILES)): %.mli.d: %.mli 712 | $(SHOW)'CAMLDEP $<' 713 | $(HIDE)$(CAMLDEP) $(OCAMLLIBS) "$<" $(redir_if_ok) 714 | 715 | $(addsuffix .d,$(ML4FILES)): %.ml4.d: %.ml4 716 | $(SHOW)'CAMLDEP -pp $<' 717 | $(HIDE)$(CAMLDEP) $(OCAMLLIBS) $(PP) -impl "$<" $(redir_if_ok) 718 | 719 | $(addsuffix .d,$(MLFILES)): %.ml.d: %.ml 720 | $(SHOW)'CAMLDEP $<' 721 | $(HIDE)$(CAMLDEP) $(OCAMLLIBS) "$<" $(redir_if_ok) 722 | 723 | $(addsuffix .d,$(MLLIBFILES)): %.mllib.d: %.mllib 724 | $(SHOW)'COQDEP $<' 725 | $(HIDE)$(COQDEP) $(OCAMLLIBS) -c "$<" $(redir_if_ok) 726 | 727 | $(addsuffix .d,$(MLPACKFILES)): %.mlpack.d: %.mlpack 728 | $(SHOW)'COQDEP $<' 729 | $(HIDE)$(COQDEP) $(OCAMLLIBS) -c "$<" $(redir_if_ok) 730 | 731 | # If this makefile is created using a _CoqProject we have coqdep get 732 | # options from it. This avoids argument length limits for pathological 733 | # projects. Note that extra options might be on the command line. 734 | VDFILE_FLAGS:=$(if _CoqProject,-f _CoqProject,) $(CMDLINE_COQLIBS) $(CMDLINE_VFILES) 735 | 736 | $(VDFILE).d: $(VFILES) 737 | $(SHOW)'COQDEP VFILES' 738 | $(HIDE)$(COQDEP) -dyndep var $(VDFILE_FLAGS) $(redir_if_ok) 739 | 740 | # Misc ######################################################################## 741 | 742 | byte: 743 | $(HIDE)$(MAKE) all "OPT:=-byte" -f "$(SELF)" 744 | .PHONY: byte 745 | 746 | opt: 747 | $(HIDE)$(MAKE) all "OPT:=-opt" -f "$(SELF)" 748 | .PHONY: opt 749 | 750 | # This is deprecated. To extend this makefile use 751 | # extension points and makefile.local 752 | printenv:: 753 | $(warning printenv is deprecated) 754 | $(warning write extensions in makefile.local or include makefile.conf) 755 | @echo 'LOCAL = $(LOCAL)' 756 | @echo 'COQLIB = $(COQLIB)' 757 | @echo 'DOCDIR = $(DOCDIR)' 758 | @echo 'OCAMLFIND = $(OCAMLFIND)' 759 | @echo 'CAMLP5O = $(CAMLP5O)' 760 | @echo 'CAMLP5BIN = $(CAMLP5BIN)' 761 | @echo 'CAMLP5LIB = $(CAMLP5LIB)' 762 | @echo 'CAMLP5OPTIONS = $(CAMLP5OPTIONS)' 763 | @echo 'HASNATDYNLINK = $(HASNATDYNLINK)' 764 | @echo 'SRC_SUBDIRS = $(SRC_SUBDIRS)' 765 | @echo 'COQ_SRC_SUBDIRS = $(COQ_SRC_SUBDIRS)' 766 | @echo 'OCAMLFIND = $(OCAMLFIND)' 767 | @echo 'PP = $(PP)' 768 | @echo 'COQFLAGS = $(COQFLAGS)' 769 | @echo 'COQLIB = $(COQLIBS)' 770 | @echo 'COQLIBINSTALL = $(COQLIBINSTALL)' 771 | @echo 'COQDOCINSTALL = $(COQDOCINSTALL)' 772 | .PHONY: printenv 773 | 774 | # Generate a .merlin file. If you need to append directives to this 775 | # file you can extend the merlin-hook target in makefile.local 776 | .merlin: 777 | $(SHOW)'FILL .merlin' 778 | $(HIDE)echo 'FLG $(COQMF_CAMLFLAGS)' > .merlin 779 | $(HIDE)echo 'B $(COQLIB)' >> .merlin 780 | $(HIDE)echo 'S $(COQLIB)' >> .merlin 781 | $(HIDE)$(foreach d,$(COQ_SRC_SUBDIRS), \ 782 | echo 'B $(COQLIB)$(d)' >> .merlin;) 783 | $(HIDE)$(foreach d,$(COQ_SRC_SUBDIRS), \ 784 | echo 'S $(COQLIB)$(d)' >> .merlin;) 785 | $(HIDE)$(foreach d,$(SRC_SUBDIRS), echo 'B $(d)' >> .merlin;) 786 | $(HIDE)$(foreach d,$(SRC_SUBDIRS), echo 'S $(d)' >> .merlin;) 787 | $(HIDE)$(MAKE) merlin-hook -f "$(SELF)" 788 | .PHONY: merlin 789 | 790 | merlin-hook:: 791 | @# Extension point 792 | .PHONY: merlin-hook 793 | 794 | # prints all variables 795 | debug: 796 | $(foreach v,\ 797 | $(sort $(filter-out $(INITIAL_VARS) INITIAL_VARS,\ 798 | $(.VARIABLES))),\ 799 | $(info $(v) = $($(v)))) 800 | .PHONY: debug 801 | 802 | .DEFAULT_GOAL := all 803 | --------------------------------------------------------------------------------