├── .gitignore ├── .travis.yml ├── Arith.v ├── LinearMemory.v ├── Makefile ├── Memory.v ├── README.md ├── Tactics.v ├── Vars.v ├── _CoqProject └── docs └── paper.pdf /.gitignore: -------------------------------------------------------------------------------- 1 | *.vo 2 | *.glob 3 | *.aux 4 | *.v.d 5 | html/ 6 | -------------------------------------------------------------------------------- /.travis.yml: -------------------------------------------------------------------------------- 1 | before_install: 2 | - travis_retry sudo add-apt-repository -y ppa:jgross-h/many-coq-versions 3 | - travis_retry sudo apt-get update 4 | - travis_retry sudo apt-get install coq-8.5pl3 5 | script: 6 | - make 7 | sudo: required 8 | -------------------------------------------------------------------------------- /Arith.v: -------------------------------------------------------------------------------- 1 | (** Calculation of the simple arithmetic language. *) 2 | 3 | Require Import Tactics. 4 | Require Export Memory. 5 | Module Arith (mem : Memory). 6 | Import mem. 7 | 8 | (** * Syntax *) 9 | 10 | Inductive Expr : Set := 11 | | Val : nat -> Expr 12 | | Add : Expr -> Expr -> Expr. 13 | 14 | (** * Semantics *) 15 | 16 | Fixpoint eval (x: Expr) : nat := 17 | match x with 18 | | Val n => n 19 | | Add x1 x2 => eval x1 + eval x2 20 | end. 21 | 22 | (** * Compiler *) 23 | 24 | Inductive Code : Set := 25 | | LOAD : nat -> Code -> Code 26 | | ADD : adr -> Code -> Code 27 | | STORE : adr -> Code -> Code 28 | | HALT : Code. 29 | 30 | Fixpoint comp' (x : Expr) (r : adr) (c : Code) : Code := 31 | match x with 32 | | Val n => LOAD n c 33 | | Add x1 x2 => comp' x1 r (STORE r (comp' x2 (next r) (ADD r c))) 34 | end. 35 | 36 | Definition comp (x : Expr) : Code := comp' x first HALT. 37 | 38 | (** * Virtual Machine *) 39 | 40 | Definition Memory : Type := Mem. 41 | Definition Acc : Type := nat. 42 | 43 | Definition State : Type := (Code * Memory * Acc)%type. 44 | 45 | Reserved Notation "x ==> y" (at level 80, no associativity). 46 | Inductive VM : State -> State -> Prop := 47 | | vm_load n c m a : (LOAD n c , m, a) ==> (c , m, n) 48 | | vm_add c r m a : (ADD r c, m, a) ==> (c, free r m, get r m + a) 49 | | vm_store c r m a : (STORE r c, m, a) ==> (c, set r a m, a) 50 | where "x ==> y" := (VM x y). 51 | 52 | (** * Calculation *) 53 | 54 | (** Boilerplate to import calculation tactics *) 55 | 56 | Module VM <: Preorder. 57 | Definition State := State. 58 | Definition VM := VM. 59 | End VM. 60 | Module VMCalc := Calculation VM. 61 | Import VMCalc. 62 | 63 | (** Specification of the compiler *) 64 | 65 | Theorem spec x r c a m : 66 | isFreeFrom r m -> 67 | (comp' x r c, m, a) =>> (c , m, eval x). 68 | 69 | (** Setup the induction proof *) 70 | 71 | Proof. 72 | intros. 73 | generalize dependent c. 74 | generalize dependent m. 75 | generalize dependent r. 76 | generalize dependent a. 77 | induction x;intros. 78 | 79 | (** Calculation of the compiler *) 80 | 81 | (** - [x = Val n]: *) 82 | 83 | begin 84 | (c, m, n). 85 | <== { apply vm_load } 86 | (LOAD n c, m, a). 87 | []. 88 | 89 | (** - [x = Add x1 x2]: *) 90 | 91 | begin 92 | (c, m, eval x1 + eval x2). 93 | = {rewrite isFreeFrom_free, get_set} 94 | (c, free r (set r (eval x1) m), get r (set r (eval x1) m) + eval x2). 95 | <== {apply vm_add} 96 | (ADD r c, set r (eval x1) m, eval x2). 97 | <<= {apply IHx2; auto using isFreeFrom_set} 98 | (comp' x2 (next r) (ADD r c), set r (eval x1) m, eval x1). 99 | <== {apply vm_store} 100 | (STORE r (comp' x2 (next r) (ADD r c)), m, eval x1). 101 | <<= { apply IHx1} 102 | (comp' x1 r (STORE r (comp' x2 (next r) (ADD r c))), m, a). 103 | []. 104 | Qed. 105 | 106 | 107 | (** * Soundness *) 108 | 109 | (** Since the VM is defined as a small step operational semantics, we 110 | have to prove that the VM is deterministic and does not get stuck in 111 | order to derive soundness from the above theorem. *) 112 | 113 | 114 | Lemma determ_vm : determ VM. 115 | intros C C1 C2 V. induction V; intro V'; inversion V'; subst; reflexivity. 116 | Qed. 117 | 118 | 119 | Theorem sound x a C : (comp x, empty, a) =>>! C -> C = (HALT , empty, eval x). 120 | Proof. 121 | intros. 122 | pose (spec x first HALT) as H'. unfold comp in *. pose (determ_trc determ_vm) as D. 123 | unfold determ in D. eapply D. apply H. split. apply H'. apply isFreeFrom_first. intro Contra. destruct Contra. 124 | inversion H0. 125 | Qed. 126 | 127 | End Arith. -------------------------------------------------------------------------------- /LinearMemory.v: -------------------------------------------------------------------------------- 1 | Require Import Memory. 2 | Require Import ZArith. 3 | Require Import FunctionalExtensionality. 4 | Module LinearMemory <: Memory. 5 | Definition adr := nat. 6 | Definition first := 0. 7 | Definition next := S. 8 | 9 | 10 | Definition Mem := adr -> nat. 11 | 12 | Definition get r (m : Mem) := m r. 13 | Definition set r v (m : Mem) r' := 14 | if beq_nat r' r then v else m r'. 15 | 16 | 17 | Definition empty : Mem := fun _ => 0. 18 | 19 | Definition free : adr -> Mem -> Mem := 20 | fun r m => set r 0 m. 21 | 22 | Definition isFreeFrom : adr -> Mem -> Prop := 23 | fun r m => forall r', r <= r' -> m r' = 0. 24 | 25 | Lemma isFreeFrom_free : forall (r : adr) (m : Mem) (n : nat), isFreeFrom r m -> free r (set r n m) = m. 26 | Proof. 27 | intros. unfold free, isFreeFrom, set in *. apply functional_extensionality. intros. 28 | remember (beq_nat x r) as b. destruct b. 29 | - symmetry in Heqb. apply beq_nat_true in Heqb. subst. erewrite H;auto. 30 | - reflexivity. 31 | Qed. 32 | 33 | Lemma isFreeFrom_set : forall (r : adr) (m : Mem) (n : nat), isFreeFrom r m -> isFreeFrom (next r) (set r n m). 34 | Proof. 35 | unfold isFreeFrom, set, next. intros. remember (beq_nat r' r) as b. destruct b;symmetry in Heqb. 36 | - apply beq_nat_true in Heqb. omega. 37 | - apply beq_nat_false in Heqb. apply H. omega. 38 | Qed. 39 | 40 | Lemma get_set : forall (r : adr) (m : Mem) (n : nat), get r (set r n m) = n. 41 | Proof. 42 | intros. unfold get, set. rewrite <- beq_nat_refl. reflexivity. 43 | Qed. 44 | 45 | Lemma isFreeFrom_first : isFreeFrom first empty. 46 | Proof. 47 | unfold isFreeFrom, first, empty. auto. 48 | Qed. 49 | End LinearMemory. 50 | 51 | Module LinearMemoryExt <: MemoryExt. 52 | Include LinearMemory. 53 | 54 | Lemma get_set' : forall (r : adr) (r' : adr) (m : Mem) (n : nat), 55 | r <> r' -> get r (set r' n m) = get r m. 56 | Proof. 57 | intros. 58 | unfold get, set. remember (r =? r') as b. destruct b. 59 | - symmetry in Heqb. apply beq_nat_true in Heqb. contradiction. 60 | - reflexivity. 61 | Qed. 62 | 63 | 64 | Lemma next_fresh : forall (r r' : adr), lta adr next r r' -> r <> r'. 65 | Proof. 66 | intros. apply Nat.lt_neq. 67 | intros. induction H; auto. 68 | Qed. 69 | 70 | End LinearMemoryExt. -------------------------------------------------------------------------------- /Makefile: -------------------------------------------------------------------------------- 1 | ############################################################################# 2 | ## v # The Coq Proof Assistant ## 3 | ## $@ 217 | printf 'cd "$${DSTROOT}"$(COQLIBINSTALL)/ && rm -f $(VOFILES) $(VFILES) $(GLOBFILES) $(NATIVEFILES) $(CMOFILES) $(CMIFILES) $(CMAFILES) && find . -type d -and -empty -delete\ncd "$${DSTROOT}"$(COQLIBINSTALL) && find "" -maxdepth 0 -and -empty -exec rmdir -p \{\} \;\n' >> "$@" 218 | printf 'cd "$${DSTROOT}"$(COQDOCINSTALL)/$(INSTALLDEFAULTROOT) \\\n' >> "$@" 219 | printf '&& rm -f $(shell find "html" -maxdepth 1 -and -type f -print)\n' >> "$@" 220 | printf 'cd "$${DSTROOT}"$(COQDOCINSTALL) && find $(INSTALLDEFAULTROOT)/html -maxdepth 0 -and -empty -exec rmdir -p \{\} \;\n' >> "$@" 221 | chmod +x $@ 222 | 223 | uninstall: uninstall_me.sh 224 | sh $< 225 | 226 | .merlin: 227 | @echo 'FLG -rectypes' > .merlin 228 | @echo "B $(COQLIB) kernel" >> .merlin 229 | @echo "B $(COQLIB) lib" >> .merlin 230 | @echo "B $(COQLIB) library" >> .merlin 231 | @echo "B $(COQLIB) parsing" >> .merlin 232 | @echo "B $(COQLIB) pretyping" >> .merlin 233 | @echo "B $(COQLIB) interp" >> .merlin 234 | @echo "B $(COQLIB) printing" >> .merlin 235 | @echo "B $(COQLIB) intf" >> .merlin 236 | @echo "B $(COQLIB) proofs" >> .merlin 237 | @echo "B $(COQLIB) tactics" >> .merlin 238 | @echo "B $(COQLIB) tools" >> .merlin 239 | @echo "B $(COQLIB) toplevel" >> .merlin 240 | @echo "B $(COQLIB) stm" >> .merlin 241 | @echo "B $(COQLIB) grammar" >> .merlin 242 | @echo "B $(COQLIB) config" >> .merlin 243 | 244 | clean:: 245 | rm -f $(OBJFILES) $(OBJFILES:.o=.native) $(NATIVEFILES) 246 | find . -name .coq-native -type d -empty -delete 247 | rm -f $(VOFILES) $(VOFILES:.vo=.vio) $(GFILES) $(VFILES:.v=.v.d) $(VFILES:=.beautified) $(VFILES:=.old) 248 | rm -f all.ps all-gal.ps all.pdf all-gal.pdf all.glob $(VFILES:.v=.glob) $(VFILES:.v=.tex) $(VFILES:.v=.g.tex) all-mli.tex 249 | - rm -rf html mlihtml uninstall_me.sh 250 | 251 | cleanall:: clean 252 | rm -f $(patsubst %.v,.%.aux,$(VFILES)) 253 | 254 | archclean:: 255 | rm -f *.cmx *.o 256 | 257 | printenv: 258 | @"$(COQBIN)coqtop" -config 259 | @echo 'CAMLC = $(CAMLC)' 260 | @echo 'CAMLOPTC = $(CAMLOPTC)' 261 | @echo 'PP = $(PP)' 262 | @echo 'COQFLAGS = $(COQFLAGS)' 263 | @echo 'COQLIBINSTALL = $(COQLIBINSTALL)' 264 | @echo 'COQDOCINSTALL = $(COQDOCINSTALL)' 265 | 266 | Makefile: _CoqProject 267 | mv -f $@ $@.bak 268 | "$(COQBIN)coq_makefile" -f $< -o $@ 269 | 270 | 271 | ################### 272 | # # 273 | # Implicit rules. # 274 | # # 275 | ################### 276 | 277 | $(VOFILES): %.vo: %.v 278 | $(COQC) $(COQDEBUG) $(COQFLAGS) $< 279 | 280 | $(GLOBFILES): %.glob: %.v 281 | $(COQC) $(COQDEBUG) $(COQFLAGS) $< 282 | 283 | $(VFILES:.v=.vio): %.vio: %.v 284 | $(COQC) -quick $(COQDEBUG) $(COQFLAGS) $< 285 | 286 | $(GFILES): %.g: %.v 287 | $(GALLINA) $< 288 | 289 | $(VFILES:.v=.tex): %.tex: %.v 290 | $(COQDOC) $(COQDOCFLAGS) -latex $< -o $@ 291 | 292 | $(HTMLFILES): %.html: %.v %.glob 293 | $(COQDOC) $(COQDOCFLAGS) -html $< -o $@ 294 | 295 | $(VFILES:.v=.g.tex): %.g.tex: %.v 296 | $(COQDOC) $(COQDOCFLAGS) -latex -g $< -o $@ 297 | 298 | $(GHTMLFILES): %.g.html: %.v %.glob 299 | $(COQDOC) $(COQDOCFLAGS) -html -g $< -o $@ 300 | 301 | $(addsuffix .d,$(VFILES)): %.v.d: %.v 302 | $(COQDEP) $(COQLIBS) "$<" > "$@" || ( RV=$$?; rm -f "$@"; exit $${RV} ) 303 | 304 | $(addsuffix .beautified,$(VFILES)): %.v.beautified: 305 | $(COQC) $(COQDEBUG) $(COQFLAGS) -beautify $*.v 306 | 307 | # WARNING 308 | # 309 | # This Makefile has been automagically generated 310 | # Edit at your own risks ! 311 | # 312 | # END OF WARNING 313 | 314 | -------------------------------------------------------------------------------- /Memory.v: -------------------------------------------------------------------------------- 1 | Module Type Memory. 2 | 3 | Parameter adr : Set. 4 | 5 | Parameter first : adr. 6 | Parameter next : adr -> adr. 7 | 8 | 9 | 10 | Parameter Mem : Type. 11 | 12 | Parameter empty : Mem. 13 | Parameter get : adr -> Mem -> nat. 14 | Parameter set : adr -> nat -> Mem -> Mem. 15 | Parameter free : adr -> Mem -> Mem. 16 | Parameter isFreeFrom : adr -> Mem -> Prop. 17 | 18 | Axiom isFreeFrom_free : forall (r : adr) (m : Mem) (n : nat), isFreeFrom r m -> free r (set r n m) = m. 19 | 20 | Axiom isFreeFrom_set : forall (r : adr) (m : Mem) (n : nat), isFreeFrom r m -> isFreeFrom (next r) (set r n m). 21 | 22 | Axiom get_set : forall (r : adr) (m : Mem) (n : nat), get r (set r n m) = n. 23 | 24 | Axiom isFreeFrom_first : isFreeFrom first empty. 25 | 26 | 27 | End Memory. 28 | 29 | 30 | Inductive lta T n (r : T) : T -> Prop := 31 | | lea_self : lta T n r (n r) 32 | | lea_next r' : lta T n r r' -> lta T n r (n r'). 33 | 34 | 35 | (* Extended memory model with additional axioms that are only used for 36 | compiling variables. *) 37 | 38 | Module Type MemoryExt. 39 | Include Memory. 40 | 41 | 42 | Axiom get_set' : forall (r : adr) (r' : adr) (m : Mem) (n : nat), 43 | r <> r' -> get r (set r' n m) = get r m. 44 | 45 | Infix "<" := (lta adr next). 46 | Hint Constructors lta. 47 | 48 | Axiom next_fresh : forall (r r' : adr), r < r' -> r <> r'. 49 | 50 | End MemoryExt. -------------------------------------------------------------------------------- /README.md: -------------------------------------------------------------------------------- 1 | # Compiling a Fifty Year Journey [![Build Status](https://travis-ci.org/pa-ba/McCarthy-Painter.svg?branch=master)](https://travis-ci.org/pa-ba/McCarthy-Painter) 2 | 3 | This repository contains the supplementary material for the paper 4 | ["Compiling a Fifty Year Journey"](docs/paper.pdf) 5 | by Graham Hutton and Patrick Bahr. The material includes Coq 6 | formalisations of the calculations in the paper. 7 | 8 | 9 | ## File Structure 10 | 11 | 12 | Below we list the relevant Coq files for the calculations in the 13 | paper: 14 | 15 | - [Arith.v](Arith.v): the calculations from the paper (arithmetic expressions) 16 | - [Vars.v](Vars.v): the calculations for the language from McCarthy & 17 | Painter paper (arithmetic expressions + variables) 18 | - [Memory.v](Memory.v): the abstract memory model 19 | - [LinearMemory.v](LinearMemory.v): simple instantiation of the memory model 20 | - [Tactics.v](Tactics.v): tactics library 21 | 22 | ## Paper vs. Coq Proofs 23 | 24 | 25 | The Coq proofs proceed as the calculations in the paper. There are, 26 | however, two minor technical difference due to the nature of the Coq 27 | system. 28 | 29 | 1. In the paper the derived VMs are tail recursive, first-order 30 | functions. The Coq system must be able to prove termination of 31 | all recursive function definitions. Since Coq's termination 32 | checker is not powerful enough to prove termination for some of 33 | the VMs (VMs from sections 3.1, 4.1, 5) or the VMs are not 34 | expected to terminate in general (VMs for lambda calculi / for 35 | language with loops), we had to define the VMs as relations 36 | instead. In particular, all VMs are defined as a small-step 37 | semantics. Each tail recursive function of a VM corresponds to a 38 | configuration constructor in the small-step semantics. As a 39 | consequence, the calculations do not prove equations, but rather 40 | instances of the relation `=>>`, which is the transitive, 41 | reflexive closure of the relation `==>` that defines the VM. 42 | 43 | 2. The Coq files contain the final result of the calculation, and 44 | thus do not reflect the *process* of discovering the definition 45 | of the compiler and the VM. That is, the files already contain 46 | the full definitions of the compiler and the virtual machine. But 47 | we used the same methodology as described in the paper to 48 | *develop* the Coq proofs. This is achieved by initially defining 49 | the `Code` data type as an empty type, defining the `==>` 50 | relation as an empty relation (i.e. with no rules), and defining 51 | the compiler function using the term `Admit` (which corresponds 52 | to Haskell's "undefined"). This setup then allows us to calculate 53 | the definition of the `Code` data type, the VM, and the compiler 54 | as described in the paper. 55 | 56 | 57 | ## Technical Details 58 | 59 | ### Dependencies 60 | 61 | - To check the proofs: Coq 8.5pl3 62 | - To step through the proofs: GNU Emacs 24.5.1, Proof General 4.4 63 | 64 | ### Proof Checking 65 | 66 | The complete Coq development in this repository is proof-checked 67 | automatically. The current status is: 68 | [![Build Status](https://travis-ci.org/pa-ba/McCarthy-Painter.svg?branch=master)](https://travis-ci.org/pa-ba/McCarthy-Painter) 69 | 70 | To check and compile the complete Coq development yourself, you can 71 | use the `Makefile`: 72 | 73 | ```shell 74 | > make 75 | -------------------------------------------------------------------------------- /Tactics.v: -------------------------------------------------------------------------------- 1 | Definition Admit {A} : A. admit. Admitted. 2 | 3 | Ltac rewr_assumption := idtac; match goal with 4 | | [R: _ = _ |- _ ] => first [rewrite R| rewrite <- R] 5 | end. 6 | 7 | 8 | Module Type Preorder. 9 | 10 | Parameter State : Type. 11 | Parameter VM : State -> State -> Prop. 12 | 13 | End Preorder. 14 | 15 | Module Calculation (Ord : Preorder). 16 | Import Ord. 17 | 18 | Notation "x ==> y" := (VM x y) (at level 80, no associativity). 19 | 20 | Reserved Notation "x =>> y" (at level 80, no associativity). 21 | Inductive trc : State -> State -> Prop := 22 | | trc_refl c : c =>> c 23 | | trc_step_trans c1 c2 c3 : c1 ==> c2 -> c2 =>> c3 -> c1 =>> c3 24 | where "x =>> y" := (trc x y). 25 | 26 | 27 | Lemma trc_step c1 c2 : c1 ==> c2 -> c1 =>> c2. 28 | Proof. 29 | intros. 30 | eapply trc_step_trans. eassumption. apply trc_refl. 31 | Qed. 32 | 33 | Lemma trc_trans c1 c2 c3 : c1 =>> c2 -> c2 =>> c3 -> c1 =>> c3. 34 | Proof. 35 | intros T S. 36 | induction T. assumption. eapply trc_step_trans. eassumption. apply IHT. assumption. 37 | Qed. 38 | 39 | 40 | Corollary trc_step_trans' c1 c2 c3 : c1 =>> c2 -> c2 ==> c3 -> c1 =>> c3. 41 | Proof. 42 | intros. eapply trc_trans. eassumption. apply trc_step. assumption. 43 | Qed. 44 | 45 | Corollary trc_eq_trans c1 c2 c3 : c1 =>> c2 -> c2 = c3 -> c1 =>> c3. 46 | Proof. 47 | intros. eapply trc_trans. eassumption. subst. apply trc_refl. 48 | Qed. 49 | 50 | Ltac smart_destruct x := first[is_var x;destruct x| let x' := fresh in remember x as x'; destruct x' ]. 51 | 52 | Ltac dist t := idtac; subst; simpl; try solve [t;try rewr_assumption;auto|apply trc_step;t;eauto 53 | |apply trc_refl;t;eauto] ; match goal with 54 | | [ H : ex _ |- _ ] => destruct H; dist t 55 | | [ H : or _ _ |- _ ] => destruct H; dist t 56 | | [ |- context [let _ := ?x in _] ] => smart_destruct x;dist t 57 | | [ |- context [match ?x with _ => _ end]] => smart_destruct x; dist t 58 | end. 59 | 60 | Ltac dist_refl := dist reflexivity. 61 | 62 | 63 | Ltac check_exp' x y t := let h := fresh "check" in assert (h: x = y) by t; try rewrite <- h; clear h. 64 | Ltac check_exp x y := let h := fresh "check" in assert (h: x = y) by reflexivity; clear h. 65 | 66 | Ltac check_rel R Rel := first [check_exp R Rel| 67 | fail 2 "wrong goal; expected relation" R "but found" Rel]. 68 | 69 | 70 | 71 | Tactic Notation "[]" := apply trc_refl. 72 | 73 | 74 | 75 | 76 | Ltac step rel lem t1 e2 := 77 | match goal with 78 | | [|- ?Rel ?lhs ?rhs] => check_rel trc Rel; 79 | first [let h := fresh "rewriting" in 80 | assert(h : rel e2 rhs) by (dist t1) ; apply (fun x => lem _ _ _ x h); clear h | fail 2] 81 | | _ => fail 1 "goal is not a VM" 82 | end. 83 | 84 | Tactic Notation (at level 2) "<<=" "{"tactic(t) "}" constr(e) := 85 | step trc trc_trans t e. 86 | 87 | Tactic Notation (at level 2) "=" "{"tactic(t) "}" constr(e) := 88 | step (@eq State) trc_eq_trans t e. 89 | 90 | Tactic Notation (at level 2) "<==" "{"tactic(t) "}" constr(e) := 91 | step VM trc_step_trans' t e. 92 | 93 | Ltac step_try rel e2 := 94 | match goal with 95 | | [|- ?Rel ?lhs ?rhs] => check_rel trc Rel; 96 | first [let h := fresh "step_try" in assert(h : rel e2 rhs)|fail 2] 97 | | _ => fail 1 "goal is not a VM" 98 | end. 99 | 100 | Tactic Notation (at level 2) "<<=" "{?}" constr(e) := step_try trc e. 101 | Tactic Notation (at level 2) "<==" "{?}" constr(e) := step_try VM e. 102 | Tactic Notation (at level 2) "=" "{?}" constr(e) := step_try (@eq State) e. 103 | 104 | Tactic Notation (at level 2) "<==" "{"tactic(t1) "}?" := 105 | match goal with 106 | | [|- ?Rel ?lhs ?rhs] => check_rel trc Rel; 107 | first [eapply trc_trans; [idtac|solve[t1]] | fail 2] 108 | | _ => fail 1 "goal is not a VM" 109 | end. 110 | 111 | Tactic Notation (at level 2) "begin" constr(rhs) := match goal with 112 | | [|- ?Rel ?lhs ?rhs'] => check_rel trc Rel; check_exp' rhs rhs' dist_refl 113 | | _ => fail 1 "rhs does not match" 114 | end. 115 | 116 | Definition determ {A} (R : A -> A -> Prop) : Prop := forall C C1 C2, R C C1 -> R C C2 -> C1 = C2. 117 | 118 | 119 | Definition trc' C C' := C =>> C' /\ ~ exists C'', C' ==> C''. 120 | 121 | Notation "x =>>! y" := (trc' x y) (at level 80, no associativity). 122 | 123 | 124 | Lemma determ_factor C1 C2 C3 : determ VM -> C1 ==> C2 -> C1 =>>! C3 -> C2 =>> C3. 125 | Proof. 126 | unfold determ. intros. destruct H1. 127 | destruct H1. exfalso. apply H2. eexists. eassumption. 128 | 129 | assert (c2 = C2). eapply H. apply H1. apply H0. subst. assumption. 130 | Qed. 131 | 132 | 133 | Lemma determ_trc : determ VM -> determ trc'. 134 | Proof. 135 | intros. unfold determ. intros. destruct H0. 136 | induction H0. 137 | 138 | destruct H1. destruct H0. reflexivity. exfalso. apply H2. eexists. eassumption. 139 | 140 | apply IHtrc. apply H2. split. eapply determ_factor; eassumption. destruct H1. assumption. 141 | Qed. 142 | 143 | 144 | End Calculation. -------------------------------------------------------------------------------- /Vars.v: -------------------------------------------------------------------------------- 1 | (** Calculation of the simple arithmetic language. *) 2 | 3 | Require Import Tactics. 4 | Require Export Memory. 5 | Module Arith (mem : MemoryExt). 6 | Import mem. 7 | 8 | (** * Syntax *) 9 | 10 | Inductive var := X | Y | Z. 11 | 12 | Inductive Expr : Set := 13 | | Val : nat -> Expr 14 | | Add : Expr -> Expr -> Expr 15 | | Var : var -> Expr. 16 | 17 | (** * Semantics *) 18 | 19 | Definition env := var -> nat. 20 | 21 | Fixpoint eval (e : env) (x: Expr) : nat := 22 | match x with 23 | | Val n => n 24 | | Add x1 x2 => eval e x1 + eval e x2 25 | | Var v => e v 26 | end. 27 | 28 | (** * Compiler *) 29 | 30 | Inductive Code : Set := 31 | | LOAD : nat -> Code -> Code 32 | | LDA : adr -> Code -> Code 33 | | ADD : adr -> Code -> Code 34 | | STORE : adr -> Code -> Code 35 | | HALT : Code. 36 | 37 | Fixpoint comp' (x : Expr) (vars: var -> adr) (r : adr) (c : Code) : Code := 38 | match x with 39 | | Val n => LOAD n c 40 | | Add x1 x2 => comp' x1 vars r (STORE r (comp' x2 vars (next r) (ADD r c))) 41 | | Var v => LDA (vars v) c 42 | end. 43 | 44 | Definition mkVars : var -> adr := 45 | fun v => 46 | match v with 47 | | X => first 48 | | Y => next first 49 | | Z => next (next first) 50 | end. 51 | 52 | Definition firstFresh : adr := next (next (next first)). 53 | 54 | Definition mkMem (e : env) : Mem := 55 | set (mkVars Z) (e Z) ((set (mkVars Y) (e Y) (set (mkVars X) (e X) empty))). 56 | 57 | 58 | Definition comp (x : Expr) : Code := comp' x mkVars firstFresh HALT. 59 | 60 | (** * Virtual Machine *) 61 | 62 | Definition Memory : Type := Mem. 63 | Definition Acc : Type := nat. 64 | 65 | Definition State : Type := (Code * Memory * Acc)%type. 66 | 67 | Reserved Notation "x ==> y" (at level 80, no associativity). 68 | Inductive VM : State -> State -> Prop := 69 | | vm_load n c m a : (LOAD n c , m, a) ==> (c , m, n) 70 | | vm_add c r m a : (ADD r c, m, a) ==> (c, free r m, get r m + a) 71 | | vm_store c r m a : (STORE r c, m, a) ==> (c, set r a m, a) 72 | | vm_lda c r m a : (LDA r c, m, a) ==> (c, m, get r m) 73 | where "x ==> y" := (VM x y). 74 | 75 | (** * Calculation *) 76 | 77 | (** Boilerplate to import calculation tactics *) 78 | 79 | Module VM <: Preorder. 80 | Definition State := State. 81 | Definition VM := VM. 82 | End VM. 83 | Module VMCalc := Calculation VM. 84 | Import VMCalc. 85 | 86 | Lemma env_fresh (e : env) vars r m 87 | (Fresh : forall v, vars v < r) 88 | (Env : forall v, e v = get (vars v) m) v n: 89 | e v = get (vars v) (set r n m). 90 | Proof. 91 | rewrite get_set'; auto using next_fresh. 92 | Qed. 93 | 94 | 95 | (** Specification of the compiler *) 96 | 97 | Theorem spec x r c a m vars e 98 | (Fresh : forall v, vars v < r) 99 | (Env : forall v, e v = get (vars v) m): 100 | isFreeFrom r m -> 101 | (comp' x vars r c, m, a) =>> (c , m, eval e x). 102 | 103 | (** Setup the induction proof *) 104 | 105 | Proof. 106 | intros. 107 | generalize dependent c. 108 | generalize dependent m. 109 | generalize dependent r. 110 | generalize dependent a. 111 | induction x;intros. 112 | 113 | (** Calculation of the compiler *) 114 | 115 | (** - [x = Val n]: *) 116 | 117 | begin 118 | (c, m, n). 119 | <== { apply vm_load } 120 | (LOAD n c, m, a). 121 | []. 122 | 123 | (** - [x = Add x1 x2]: *) 124 | 125 | begin 126 | (c, m, eval e x1 + eval e x2). 127 | = {rewrite isFreeFrom_free, get_set} 128 | (c, free r (set r (eval e x1) m), get r (set r (eval e x1) m) + eval e x2). 129 | <== {apply vm_add} 130 | (ADD r c, set r (eval e x1) m, eval e x2). 131 | <<= {apply IHx2; auto using isFreeFrom_set, env_fresh} 132 | (comp' x2 vars (next r) (ADD r c), set r (eval e x1) m, eval e x1). 133 | <== {apply vm_store} 134 | (STORE r (comp' x2 vars (next r) (ADD r c)), m, eval e x1). 135 | <<= { apply IHx1} 136 | (comp' x1 vars r (STORE r (comp' x2 vars (next r) (ADD r c))), m, a). 137 | []. 138 | 139 | (** - [x = Var v]: *) 140 | 141 | begin 142 | (c, m, e v). 143 | = {rewrite Env} 144 | (c, m, get (vars v) m). 145 | <<= {apply vm_lda} 146 | (LDA (vars v) c , m, a). 147 | []. 148 | Qed. 149 | 150 | 151 | Lemma mkMemEnv (e : env) (v : var) : e v = get (mkVars v) (mkMem e). 152 | Proof. 153 | destruct v; 154 | unfold mkMem; simpl; 155 | repeat first [rewrite get_set' by (apply next_fresh; auto) | rewrite get_set]; 156 | reflexivity. 157 | Qed. 158 | 159 | Lemma mkVarsFresh v : mkVars v < firstFresh. 160 | Proof. 161 | unfold firstFresh. destruct v; simpl; auto. 162 | Qed. 163 | 164 | Lemma isFreeFromFirst e : isFreeFrom firstFresh (mkMem e). 165 | Proof. 166 | unfold mkMem. simpl. eauto using isFreeFrom_first, isFreeFrom_set. 167 | Qed. 168 | 169 | Theorem spec_top x a e: 170 | (comp x, mkMem e, a) =>> (HALT , mkMem e, eval e x). 171 | Proof. 172 | apply spec. 173 | - apply mkVarsFresh. 174 | - apply mkMemEnv. 175 | - apply isFreeFromFirst. 176 | Qed. 177 | 178 | 179 | (** * Soundness *) 180 | 181 | (** Since the VM is defined as a small step operational semantics, we 182 | have to prove that the VM is deterministic and does not get stuck in 183 | order to derive soundness from the above theorem. *) 184 | 185 | 186 | Lemma determ_vm : determ VM. 187 | intros C C1 C2 V. induction V; intro V'; inversion V'; subst; reflexivity. 188 | Qed. 189 | 190 | 191 | Theorem sound x a C e : (comp x, mkMem e, a) =>>! C -> C = (HALT , mkMem e, eval e x). 192 | Proof. 193 | intros. 194 | pose (spec_top x) as H'. unfold comp in *. pose (determ_trc determ_vm) as D. 195 | unfold determ in D. eapply D. apply H. split. apply H'. intro Contra. destruct Contra. 196 | inversion H0. 197 | Qed. 198 | 199 | End Arith. -------------------------------------------------------------------------------- /_CoqProject: -------------------------------------------------------------------------------- 1 | -R . "" Memory.v 2 | -R . "" LinearMemory.v 3 | -R . "" Tactics.v 4 | -R . "" Arith.v 5 | -R . "" Vars.v 6 | -------------------------------------------------------------------------------- /docs/paper.pdf: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/pa-ba/McCarthy-Painter/0ec03dd95afe5c50657a0107ed247ea2374d182c/docs/paper.pdf --------------------------------------------------------------------------------