From a7c6cf8d4d59e29ed73f9eb8bfc015b0384672a1 Mon Sep 17 00:00:00 2001 From: xleroy Date: Mon, 4 Sep 2006 15:33:05 +0000 Subject: Revu la repartition des sources Coq en sous-repertoires git-svn-id: https://yquem.inria.fr/compcert/svn/compcert/trunk@73 fca1b0fc-160b-0410-b1d3-a4f43f01ea2e --- backend/AST.v | 250 ----- backend/Cminorgen.v | 433 -------- backend/Cminorgenproof.v | 2595 ---------------------------------------------- backend/Csharpminor.v | 563 ---------- backend/Events.v | 103 -- backend/Globalenvs.v | 643 ------------ backend/Main.v | 311 ------ backend/Mem.v | 2385 ------------------------------------------ backend/Values.v | 888 ---------------- 9 files changed, 8171 deletions(-) delete mode 100644 backend/AST.v delete mode 100644 backend/Cminorgen.v delete mode 100644 backend/Cminorgenproof.v delete mode 100644 backend/Csharpminor.v delete mode 100644 backend/Events.v delete mode 100644 backend/Globalenvs.v delete mode 100644 backend/Main.v delete mode 100644 backend/Mem.v delete mode 100644 backend/Values.v (limited to 'backend') diff --git a/backend/AST.v b/backend/AST.v deleted file mode 100644 index 1342bef1..00000000 --- a/backend/AST.v +++ /dev/null @@ -1,250 +0,0 @@ -(** This file defines a number of data types and operations used in - the abstract syntax trees of many of the intermediate languages. *) - -Require Import Coqlib. -Require Import Integers. -Require Import Floats. - -Set Implicit Arguments. - -(** Identifiers (names of local variables, of global symbols and functions, - etc) are represented by the type [positive] of positive integers. *) - -Definition ident := positive. - -Definition ident_eq := peq. - -(** The languages are weakly typed, using only two types: [Tint] for - integers and pointers, and [Tfloat] for floating-point numbers. *) - -Inductive typ : Set := - | Tint : typ - | Tfloat : typ. - -Definition typesize (ty: typ) : Z := - match ty with Tint => 4 | Tfloat => 8 end. - -(** Additionally, function definitions and function calls are annotated - by function signatures indicating the number and types of arguments, - as well as the type of the returned value if any. These signatures - are used in particular to determine appropriate calling conventions - for the function. *) - -Record signature : Set := mksignature { - sig_args: list typ; - sig_res: option typ -}. - -Definition proj_sig_res (s: signature) : typ := - match s.(sig_res) with - | None => Tint - | Some t => t - end. - -(** Memory accesses (load and store instructions) are annotated by - a ``memory chunk'' indicating the type, size and signedness of the - chunk of memory being accessed. *) - -Inductive memory_chunk : Set := - | Mint8signed : memory_chunk (**r 8-bit signed integer *) - | Mint8unsigned : memory_chunk (**r 8-bit unsigned integer *) - | Mint16signed : memory_chunk (**r 16-bit signed integer *) - | Mint16unsigned : memory_chunk (**r 16-bit unsigned integer *) - | Mint32 : memory_chunk (**r 32-bit integer, or pointer *) - | Mfloat32 : memory_chunk (**r 32-bit single-precision float *) - | Mfloat64 : memory_chunk. (**r 64-bit double-precision float *) - -(** Initialization data for global variables. *) - -Inductive init_data: Set := - | Init_int8: int -> init_data - | Init_int16: int -> init_data - | Init_int32: int -> init_data - | Init_float32: float -> init_data - | Init_float64: float -> init_data - | Init_space: Z -> init_data. - -(** Whole programs consist of: -- a collection of function definitions (name and description); -- the name of the ``main'' function that serves as entry point in the program; -- a collection of global variable declarations (name and initializer). - -The type of function descriptions varies among the various intermediate -languages and is taken as parameter to the [program] type. The other parts -of whole programs are common to all languages. *) - -Record program (funct: Set) : Set := mkprogram { - prog_funct: list (ident * funct); - prog_main: ident; - prog_vars: list (ident * list init_data) -}. - -(** We now define a general iterator over programs that applies a given - code transformation function to all function descriptions and leaves - the other parts of the program unchanged. *) - -Section TRANSF_PROGRAM. - -Variable A B: Set. -Variable transf: A -> B. - -Fixpoint transf_program (l: list (ident * A)) : list (ident * B) := - match l with - | nil => nil - | (id, fn) :: rem => (id, transf fn) :: transf_program rem - end. - -Definition transform_program (p: program A) : program B := - mkprogram - (transf_program p.(prog_funct)) - p.(prog_main) - p.(prog_vars). - -Remark transf_program_functions: - forall fl i tf, - In (i, tf) (transf_program fl) -> - exists f, In (i, f) fl /\ transf f = tf. -Proof. - induction fl; simpl. - tauto. - destruct a. simpl. intros. - elim H; intro. exists a. split. left. congruence. congruence. - generalize (IHfl _ _ H0). intros [f [IN TR]]. - exists f. split. right. auto. auto. -Qed. - -Lemma transform_program_function: - forall p i tf, - In (i, tf) (transform_program p).(prog_funct) -> - exists f, In (i, f) p.(prog_funct) /\ transf f = tf. -Proof. - simpl. intros. eapply transf_program_functions; eauto. -Qed. - -End TRANSF_PROGRAM. - -(** The following is a variant of [transform_program] where the - code transformation function can fail and therefore returns an - option type. *) - -Section TRANSF_PARTIAL_PROGRAM. - -Variable A B: Set. -Variable transf_partial: A -> option B. - -Fixpoint transf_partial_program - (l: list (ident * A)) : option (list (ident * B)) := - match l with - | nil => Some nil - | (id, fn) :: rem => - match transf_partial fn with - | None => None - | Some fn' => - match transf_partial_program rem with - | None => None - | Some res => Some ((id, fn') :: res) - end - end - end. - -Definition transform_partial_program (p: program A) : option (program B) := - match transf_partial_program p.(prog_funct) with - | None => None - | Some fl => Some (mkprogram fl p.(prog_main) p.(prog_vars)) - end. - -Remark transf_partial_program_functions: - forall fl tfl i tf, - transf_partial_program fl = Some tfl -> - In (i, tf) tfl -> - exists f, In (i, f) fl /\ transf_partial f = Some tf. -Proof. - induction fl; simpl. - intros; injection H; intro; subst tfl; contradiction. - case a; intros id fn. intros until tf. - caseEq (transf_partial fn). - intros tfn TFN. - caseEq (transf_partial_program fl). - intros tfl1 TFL1 EQ. injection EQ; intro; clear EQ; subst tfl. - simpl. intros [EQ1|IN1]. - exists fn. intuition congruence. - generalize (IHfl _ _ _ TFL1 IN1). - intros [f [X Y]]. - exists f. intuition congruence. - intros; discriminate. - intros; discriminate. -Qed. - -Lemma transform_partial_program_function: - forall p tp i tf, - transform_partial_program p = Some tp -> - In (i, tf) tp.(prog_funct) -> - exists f, In (i, f) p.(prog_funct) /\ transf_partial f = Some tf. -Proof. - intros until tf. - unfold transform_partial_program. - caseEq (transf_partial_program (prog_funct p)). - intros. apply transf_partial_program_functions with l; auto. - injection H0; intros; subst tp. exact H1. - intros; discriminate. -Qed. - -Lemma transform_partial_program_main: - forall p tp, - transform_partial_program p = Some tp -> - tp.(prog_main) = p.(prog_main). -Proof. - intros p tp. unfold transform_partial_program. - destruct (transf_partial_program (prog_funct p)). - intro EQ; injection EQ; intro EQ1; rewrite <- EQ1; reflexivity. - intro; discriminate. -Qed. - -End TRANSF_PARTIAL_PROGRAM. - -(** For most languages, the functions composing the program are either - internal functions, defined within the language, or external functions - (a.k.a. system calls) that emit an event when applied. We define - a type for such functions and some generic transformation functions. *) - -Record external_function : Set := mkextfun { - ef_id: ident; - ef_sig: signature -}. - -Inductive fundef (F: Set): Set := - | Internal: F -> fundef F - | External: external_function -> fundef F. - -Implicit Arguments External [F]. - -Section TRANSF_FUNDEF. - -Variable A B: Set. -Variable transf: A -> B. - -Definition transf_fundef (fd: fundef A): fundef B := - match fd with - | Internal f => Internal (transf f) - | External ef => External ef - end. - -End TRANSF_FUNDEF. - -Section TRANSF_PARTIAL_FUNDEF. - -Variable A B: Set. -Variable transf_partial: A -> option B. - -Definition transf_partial_fundef (fd: fundef A): option (fundef B) := - match fd with - | Internal f => - match transf_partial f with - | None => None - | Some f' => Some (Internal f') - end - | External ef => Some (External ef) - end. - -End TRANSF_PARTIAL_FUNDEF. - diff --git a/backend/Cminorgen.v b/backend/Cminorgen.v deleted file mode 100644 index 4c611b44..00000000 --- a/backend/Cminorgen.v +++ /dev/null @@ -1,433 +0,0 @@ -(** Translation from Csharpminor to Cminor. *) - -Require Import Coqlib. -Require Import Maps. -Require Import Sets. -Require Import AST. -Require Import Integers. -Require Mem. -Require Import Csharpminor. -Require Import Op. -Require Import Cminor. -Require Cmconstr. - -(** The main task in translating Csharpminor to Cminor is to explicitly - stack-allocate local variables whose address is taken: these local - variables become sub-blocks of the Cminor stack data block for the - current function. Taking the address of such a variable becomes - a [Oaddrstack] operation with the corresponding offset. Accessing - or assigning such a variable becomes a load or store operation at - that address. Only scalar local variables whose address is never - taken in the Csharpminor code can be mapped to Cminor local - variable, since the latter do not reside in memory. - - Other tasks performed during the translation to Cminor: -- Transformation of Csharpminor's standard set of operators and - trivial addressing modes to Cminor's processor-dependent operators - and addressing modes. This is done using the optimizing Cminor - constructor functions provided in file [Cmconstr], therefore performing - instruction selection on the fly. -- Insertion of truncation, zero- and sign-extension operations when - assigning to a Csharpminor local variable of ``small'' type - (e.g. [Mfloat32] or [Mint8signed]). This is necessary to preserve - the ``normalize at assignment-time'' semantics of Csharpminor. -*) - -(** Translation of operators. *) - -Definition make_op (op: Csharpminor.operation) (el: exprlist): option expr := - match el with - | Enil => - match op with - | Csharpminor.Ointconst n => Some(Eop (Ointconst n) Enil) - | Csharpminor.Ofloatconst n => Some(Eop (Ofloatconst n) Enil) - | _ => None - end - | Econs e1 Enil => - match op with - | Csharpminor.Ocast8unsigned => Some(Cmconstr.cast8unsigned e1) - | Csharpminor.Ocast8signed => Some(Cmconstr.cast8signed e1) - | Csharpminor.Ocast16unsigned => Some(Cmconstr.cast16unsigned e1) - | Csharpminor.Ocast16signed => Some(Cmconstr.cast16signed e1) - | Csharpminor.Onotint => Some(Cmconstr.notint e1) - | Csharpminor.Onegf => Some(Cmconstr.negfloat e1) - | Csharpminor.Oabsf => Some(Cmconstr.absfloat e1) - | Csharpminor.Osingleoffloat => Some(Cmconstr.singleoffloat e1) - | Csharpminor.Ointoffloat => Some(Cmconstr.intoffloat e1) - | Csharpminor.Ofloatofint => Some(Cmconstr.floatofint e1) - | Csharpminor.Ofloatofintu => Some(Cmconstr.floatofintu e1) - | _ => None - end - | Econs e1 (Econs e2 Enil) => - match op with - | Csharpminor.Oadd => Some(Cmconstr.add e1 e2) - | Csharpminor.Osub => Some(Cmconstr.sub e1 e2) - | Csharpminor.Omul => Some(Cmconstr.mul e1 e2) - | Csharpminor.Odiv => Some(Cmconstr.divs e1 e2) - | Csharpminor.Odivu => Some(Cmconstr.divu e1 e2) - | Csharpminor.Omod => Some(Cmconstr.mods e1 e2) - | Csharpminor.Omodu => Some(Cmconstr.modu e1 e2) - | Csharpminor.Oand => Some(Cmconstr.and e1 e2) - | Csharpminor.Oor => Some(Cmconstr.or e1 e2) - | Csharpminor.Oxor => Some(Cmconstr.xor e1 e2) - | Csharpminor.Oshl => Some(Cmconstr.shl e1 e2) - | Csharpminor.Oshr => Some(Cmconstr.shr e1 e2) - | Csharpminor.Oshru => Some(Cmconstr.shru e1 e2) - | Csharpminor.Oaddf => Some(Cmconstr.addf e1 e2) - | Csharpminor.Osubf => Some(Cmconstr.subf e1 e2) - | Csharpminor.Omulf => Some(Cmconstr.mulf e1 e2) - | Csharpminor.Odivf => Some(Cmconstr.divf e1 e2) - | Csharpminor.Ocmp c => Some(Cmconstr.cmp c e1 e2) - | Csharpminor.Ocmpu c => Some(Cmconstr.cmpu c e1 e2) - | Csharpminor.Ocmpf c => Some(Cmconstr.cmpf c e1 e2) - | _ => None - end - | _ => None - end. - -(** [make_cast chunk e] returns a Cminor expression that normalizes - the value of Cminor expression [e] as prescribed by the memory chunk - [chunk]. For instance, 8-bit sign extension is performed if - [chunk] is [Mint8signed]. *) - -Definition make_cast (chunk: memory_chunk) (e: expr): expr := - match chunk with - | Mint8signed => Cmconstr.cast8signed e - | Mint8unsigned => Cmconstr.cast8unsigned e - | Mint16signed => Cmconstr.cast16signed e - | Mint16unsigned => Cmconstr.cast16unsigned e - | Mint32 => e - | Mfloat32 => Cmconstr.singleoffloat e - | Mfloat64 => e - end. - -Definition make_load (chunk: memory_chunk) (e: expr): expr := - Cmconstr.load chunk e. - -Definition store_arg (chunk: memory_chunk) (e: expr) : expr := - match e with - | Eop op (Econs e1 Enil) => - match op with - | Ocast8signed => - match chunk with Mint8signed => e1 | _ => e end - | Ocast8unsigned => - match chunk with Mint8unsigned => e1 | _ => e end - | Ocast16signed => - match chunk with Mint16signed => e1 | _ => e end - | Ocast16unsigned => - match chunk with Mint16unsigned => e1 | _ => e end - | Osingleoffloat => - match chunk with Mfloat32 => e1 | _ => e end - | _ => e - end - | _ => e - end. - -Definition make_store (chunk: memory_chunk) (e1 e2: expr): stmt := - Sexpr(Cmconstr.store chunk e1 (store_arg chunk e2)). - -Definition make_stackaddr (ofs: Z): expr := - Eop (Oaddrstack (Int.repr ofs)) Enil. - -Definition make_globaladdr (id: ident): expr := - Eop (Oaddrsymbol id Int.zero) Enil. - -(** Compile-time information attached to each Csharpminor - variable: global variables, local variables, function parameters. - [Var_local] denotes a scalar local variable whose address is not - taken; it will be translated to a Cminor local variable of the - same name. [Var_stack_scalar] and [Var_stack_array] denote - local variables that are stored as sub-blocks of the Cminor stack - data block. [Var_global_scalar] and [Var_global_array] denote - global variables, stored in the global symbols with the same names. *) - -Inductive var_info: Set := - | Var_local: memory_chunk -> var_info - | Var_stack_scalar: memory_chunk -> Z -> var_info - | Var_stack_array: Z -> var_info - | Var_global_scalar: memory_chunk -> var_info - | Var_global_array: var_info. - -Definition compilenv := PMap.t var_info. - -(** Generation of Cminor code corresponding to accesses to Csharpminor - local variables: reads, assignments, and taking the address of. *) - -Definition var_get (cenv: compilenv) (id: ident): option expr := - match PMap.get id cenv with - | Var_local chunk => - Some(Evar id) - | Var_stack_scalar chunk ofs => - Some(make_load chunk (make_stackaddr ofs)) - | Var_global_scalar chunk => - Some(make_load chunk (make_globaladdr id)) - | _ => - None - end. - -Definition var_set (cenv: compilenv) (id: ident) (rhs: expr): option stmt := - match PMap.get id cenv with - | Var_local chunk => - Some(Sexpr(Eassign id (make_cast chunk rhs))) - | Var_stack_scalar chunk ofs => - Some(make_store chunk (make_stackaddr ofs) rhs) - | Var_global_scalar chunk => - Some(make_store chunk (make_globaladdr id) rhs) - | _ => - None - end. - -Definition var_addr (cenv: compilenv) (id: ident): option expr := - match PMap.get id cenv with - | Var_local chunk => None - | Var_stack_scalar chunk ofs => Some (make_stackaddr ofs) - | Var_stack_array ofs => Some (make_stackaddr ofs) - | _ => Some (make_globaladdr id) - end. - -(** The translation from Csharpminor to Cminor can fail, which we - encode by returning option types ([None] denotes error). - Propagation of errors is done in monadic style, using the following - [bind] monadic composition operator, and a [do] notation inspired - by Haskell's. *) - -Definition bind (A B: Set) (a: option A) (b: A -> option B): option B := - match a with - | None => None - | Some x => b x - end. - -Notation "'do' X <- A ; B" := (bind _ _ A (fun X => B)) - (at level 200, X ident, A at level 100, B at level 200). - -(** Translation of expressions. All the hard work is done by the - [make_*] and [var_*] functions defined above. *) - -Fixpoint transl_expr (cenv: compilenv) (e: Csharpminor.expr) - {struct e}: option expr := - match e with - | Csharpminor.Evar id => var_get cenv id - | Csharpminor.Eaddrof id => var_addr cenv id - | Csharpminor.Eop op el => - do tel <- transl_exprlist cenv el; make_op op tel - | Csharpminor.Eload chunk e => - do te <- transl_expr cenv e; Some (make_load chunk te) - | Csharpminor.Ecall sig e el => - do te <- transl_expr cenv e; - do tel <- transl_exprlist cenv el; - Some (Ecall sig te tel) - | Csharpminor.Econdition e1 e2 e3 => - do te1 <- transl_expr cenv e1; - do te2 <- transl_expr cenv e2; - do te3 <- transl_expr cenv e3; - Some (Cmconstr.conditionalexpr te1 te2 te3) - | Csharpminor.Elet e1 e2 => - do te1 <- transl_expr cenv e1; - do te2 <- transl_expr cenv e2; - Some (Elet te1 te2) - | Csharpminor.Eletvar n => - Some (Eletvar n) - | Csharpminor.Ealloc e => - do te <- transl_expr cenv e; - Some (Ealloc te) - end - -with transl_exprlist (cenv: compilenv) (el: Csharpminor.exprlist) - {struct el}: option exprlist := - match el with - | Csharpminor.Enil => - Some Enil - | Csharpminor.Econs e1 e2 => - do te1 <- transl_expr cenv e1; - do te2 <- transl_exprlist cenv e2; - Some (Econs te1 te2) - end. - -(** Translation of statements. Entirely straightforward. *) - -Fixpoint transl_stmt (cenv: compilenv) (s: Csharpminor.stmt) - {struct s}: option stmt := - match s with - | Csharpminor.Sskip => - Some Sskip - | Csharpminor.Sexpr e => - do te <- transl_expr cenv e; Some(Sexpr te) - | Csharpminor.Sassign id e => - do te <- transl_expr cenv e; var_set cenv id te - | Csharpminor.Sstore chunk e1 e2 => - do te1 <- transl_expr cenv e1; - do te2 <- transl_expr cenv e2; - Some (make_store chunk te1 te2) - | Csharpminor.Sseq s1 s2 => - do ts1 <- transl_stmt cenv s1; - do ts2 <- transl_stmt cenv s2; - Some (Sseq ts1 ts2) - | Csharpminor.Sifthenelse e s1 s2 => - do te <- transl_expr cenv e; - do ts1 <- transl_stmt cenv s1; - do ts2 <- transl_stmt cenv s2; - Some (Cmconstr.ifthenelse te ts1 ts2) - | Csharpminor.Sloop s => - do ts <- transl_stmt cenv s; - Some (Sloop ts) - | Csharpminor.Sblock s => - do ts <- transl_stmt cenv s; - Some (Sblock ts) - | Csharpminor.Sexit n => - Some (Sexit n) - | Csharpminor.Sswitch e cases default => - do te <- transl_expr cenv e; Some(Sswitch te cases default) - | Csharpminor.Sreturn None => - Some (Sreturn None) - | Csharpminor.Sreturn (Some e) => - do te <- transl_expr cenv e; Some (Sreturn (Some te)) - end. - -(** Computation of the set of variables whose address is taken in - a piece of Csharpminor code. *) - -Module Identset := MakeSet(PTree). - -Fixpoint addr_taken_expr (e: Csharpminor.expr): Identset.t := - match e with - | Csharpminor.Evar id => Identset.empty - | Csharpminor.Eaddrof id => Identset.add id Identset.empty - | Csharpminor.Eop op el => addr_taken_exprlist el - | Csharpminor.Eload chunk e => addr_taken_expr e - | Csharpminor.Ecall sig e el => - Identset.union (addr_taken_expr e) (addr_taken_exprlist el) - | Csharpminor.Econdition e1 e2 e3 => - Identset.union (addr_taken_expr e1) - (Identset.union (addr_taken_expr e2) (addr_taken_expr e3)) - | Csharpminor.Elet e1 e2 => - Identset.union (addr_taken_expr e1) (addr_taken_expr e2) - | Csharpminor.Eletvar n => Identset.empty - | Csharpminor.Ealloc e => addr_taken_expr e - end - -with addr_taken_exprlist (e: Csharpminor.exprlist): Identset.t := - match e with - | Csharpminor.Enil => Identset.empty - | Csharpminor.Econs e1 e2 => - Identset.union (addr_taken_expr e1) (addr_taken_exprlist e2) - end. - -Fixpoint addr_taken_stmt (s: Csharpminor.stmt): Identset.t := - match s with - | Csharpminor.Sskip => Identset.empty - | Csharpminor.Sexpr e => addr_taken_expr e - | Csharpminor.Sassign id e => addr_taken_expr e - | Csharpminor.Sstore chunk e1 e2 => - Identset.union (addr_taken_expr e1) (addr_taken_expr e2) - | Csharpminor.Sseq s1 s2 => - Identset.union (addr_taken_stmt s1) (addr_taken_stmt s2) - | Csharpminor.Sifthenelse e s1 s2 => - Identset.union (addr_taken_expr e) - (Identset.union (addr_taken_stmt s1) (addr_taken_stmt s2)) - | Csharpminor.Sloop s => addr_taken_stmt s - | Csharpminor.Sblock s => addr_taken_stmt s - | Csharpminor.Sexit n => Identset.empty - | Csharpminor.Sswitch e cases default => addr_taken_expr e - | Csharpminor.Sreturn None => Identset.empty - | Csharpminor.Sreturn (Some e) => addr_taken_expr e - end. - -(** Layout of the Cminor stack data block and construction of the - compilation environment. Csharpminor local variables that are - arrays or whose address is taken are allocated a slot in the Cminor - stack data. While this is not important for correctness, sufficient - padding is inserted to ensure adequate alignment of addresses. *) - -Definition assign_variable - (atk: Identset.t) - (id_lv: ident * var_kind) - (cenv_stacksize: compilenv * Z) : compilenv * Z := - let (cenv, stacksize) := cenv_stacksize in - match id_lv with - | (id, Varray sz) => - let ofs := align stacksize 8 in - (PMap.set id (Var_stack_array ofs) cenv, ofs + Zmax 0 sz) - | (id, Vscalar chunk) => - if Identset.mem id atk then - let sz := Mem.size_chunk chunk in - let ofs := align stacksize sz in - (PMap.set id (Var_stack_scalar chunk ofs) cenv, ofs + sz) - else - (PMap.set id (Var_local chunk) cenv, stacksize) - end. - -Fixpoint assign_variables - (atk: Identset.t) - (id_lv_list: list (ident * var_kind)) - (cenv_stacksize: compilenv * Z) - {struct id_lv_list}: compilenv * Z := - match id_lv_list with - | nil => cenv_stacksize - | id_lv :: rem => - assign_variables atk rem (assign_variable atk id_lv cenv_stacksize) - end. - -Definition build_compilenv - (globenv: compilenv) (f: Csharpminor.function) : compilenv * Z := - assign_variables - (addr_taken_stmt f.(Csharpminor.fn_body)) - (fn_variables f) - (globenv, 0). - -Definition assign_global_variable - (ce: compilenv) (info: ident * var_kind * list init_data) : compilenv := - match info with - | (id, Vscalar chunk, _) => PMap.set id (Var_global_scalar chunk) ce - | (id, Varray _, _) => PMap.set id Var_global_array ce - end. - -Definition build_global_compilenv (p: Csharpminor.program) : compilenv := - List.fold_left assign_global_variable - p.(prog_vars) (PMap.init Var_global_array). - -(** Function parameters whose address is taken must be stored in their - stack slots at function entry. (Cminor passes these parameters in - local variables.) *) - -Fixpoint store_parameters - (cenv: compilenv) (params: list (ident * memory_chunk)) - {struct params} : stmt := - match params with - | nil => Sskip - | (id, chunk) :: rem => - match PMap.get id cenv with - | Var_local chunk => - Sseq (Sexpr (Eassign id (make_cast chunk (Evar id)))) - (store_parameters cenv rem) - | Var_stack_scalar chunk ofs => - Sseq (make_store chunk (make_stackaddr ofs) (Evar id)) - (store_parameters cenv rem) - | _ => - Sskip (* should never happen *) - end - end. - -(** Translation of a Csharpminor function. We must check that the - required Cminor stack block is no bigger than [Int.max_signed], - otherwise address computations within the stack block could - overflow machine arithmetic and lead to incorrect code. *) - -Definition transl_function - (gce: compilenv) (f: Csharpminor.function): option function := - let (cenv, stacksize) := build_compilenv gce f in - if zle stacksize Int.max_signed then - do tbody <- transl_stmt cenv f.(Csharpminor.fn_body); - Some (mkfunction - (Csharpminor.fn_sig f) - (Csharpminor.fn_params_names f) - (Csharpminor.fn_vars_names f) - stacksize - (Sseq (store_parameters cenv f.(Csharpminor.fn_params)) tbody)) - else None. - -Definition transl_fundef (gce: compilenv) (f: Csharpminor.fundef): option fundef := - transf_partial_fundef (transl_function gce) f. - -Definition transl_program (p: Csharpminor.program) : option program := - let gce := build_global_compilenv p in - transform_partial_program (transl_fundef gce) (program_of_program p). diff --git a/backend/Cminorgenproof.v b/backend/Cminorgenproof.v deleted file mode 100644 index 7820095a..00000000 --- a/backend/Cminorgenproof.v +++ /dev/null @@ -1,2595 +0,0 @@ -(** Correctness proof for Cminor generation. *) - -Require Import Coqlib. -Require Import Maps. -Require Import Sets. -Require Import AST. -Require Import Integers. -Require Import Floats. -Require Import Values. -Require Import Mem. -Require Import Events. -Require Import Globalenvs. -Require Import Csharpminor. -Require Import Op. -Require Import Cminor. -Require Cmconstr. -Require Import Cminorgen. -Require Import Cmconstrproof. - -Section TRANSLATION. - -Variable prog: Csharpminor.program. -Variable tprog: program. -Hypothesis TRANSL: transl_program prog = Some tprog. -Let ge : Csharpminor.genv := Genv.globalenv (program_of_program prog). -Let tge: genv := Genv.globalenv tprog. -Let gce : compilenv := build_global_compilenv prog. - -Lemma symbols_preserved: - forall (s: ident), Genv.find_symbol tge s = Genv.find_symbol ge s. -Proof. - intro. unfold ge, tge. - apply Genv.find_symbol_transf_partial with (transl_fundef gce). - exact TRANSL. -Qed. - -Lemma function_ptr_translated: - forall (b: block) (f: Csharpminor.fundef), - Genv.find_funct_ptr ge b = Some f -> - exists tf, - Genv.find_funct_ptr tge b = Some tf /\ transl_fundef gce f = Some tf. -Proof. - intros. - generalize - (Genv.find_funct_ptr_transf_partial (transl_fundef gce) TRANSL H). - case (transl_fundef gce f). - intros tf [A B]. exists tf. tauto. - intros [A B]. elim B. reflexivity. -Qed. - -Lemma functions_translated: - forall (v: val) (f: Csharpminor.fundef), - Genv.find_funct ge v = Some f -> - exists tf, - Genv.find_funct tge v = Some tf /\ transl_fundef gce f = Some tf. -Proof. - intros. - generalize - (Genv.find_funct_transf_partial (transl_fundef gce) TRANSL H). - case (transl_fundef gce f). - intros tf [A B]. exists tf. tauto. - intros [A B]. elim B. reflexivity. -Qed. - -Lemma sig_preserved: - forall f tf, - transl_fundef gce f = Some tf -> - Cminor.funsig tf = Csharpminor.funsig f. -Proof. - intros until tf; destruct f; simpl. - unfold transl_function. destruct (build_compilenv gce f). - case (zle z Int.max_signed); try congruence. - intro. case (transl_stmt c (Csharpminor.fn_body f)); simpl; try congruence. - intros. inversion H. reflexivity. - intro. inversion H; reflexivity. -Qed. - -Definition global_compilenv_match (ce: compilenv) (gv: gvarenv) : Prop := - forall id, - match ce!!id with - | Var_global_scalar chunk => gv!id = Some (Vscalar chunk) - | Var_global_array => True - | _ => False - end. - -Lemma global_compilenv_charact: - global_compilenv_match gce (global_var_env prog). -Proof. - set (mkgve := fun gv (vars: list (ident * var_kind * list init_data)) => - List.fold_left - (fun gve x => match x with (id, k, init) => PTree.set id k gve end) - vars gv). - assert (forall vars gv ce, - global_compilenv_match ce gv -> - global_compilenv_match (List.fold_left assign_global_variable vars ce) - (mkgve gv vars)). - induction vars; simpl; intros. - auto. - apply IHvars. intro id1. unfold assign_global_variable. - destruct a as [[id2 lv2] init2]. destruct lv2; simpl; rewrite PMap.gsspec; rewrite PTree.gsspec. - case (peq id1 id2); intro. auto. apply H. - case (peq id1 id2); intro. auto. apply H. - - change (global_var_env prog) with (mkgve (PTree.empty var_kind) prog.(Csharpminor.prog_vars)). - unfold gce, build_global_compilenv. apply H. - intro. rewrite PMap.gi. auto. -Qed. - -(** * Correspondence between Csharpminor's and Cminor's environments and memory states *) - -(** In Csharpminor, every variable is stored in a separate memory block. - In the corresponding Cminor code, some of these variables reside in - the local variable environment; others are sub-blocks of the stack data - block. We capture these changes in memory via a memory injection [f]: -- [f b = None] means that the Csharpminor block [b] no longer exist - in the execution of the generated Cminor code. This corresponds to - a Csharpminor local variable translated to a Cminor local variable. -- [f b = Some(b', ofs)] means that Csharpminor block [b] corresponds - to a sub-block of Cminor block [b] at offset [ofs]. - - A memory injection [f] defines a relation [val_inject f] between - values and a relation [mem_inject f] between memory states. - These relations, defined in file [Memory], will be used intensively - in our proof of simulation between Csharpminor and Cminor executions. - - In the remainder of this section, we define relations between - Csharpminor and Cminor environments and call stacks. *) - -(** Matching for a Csharpminor variable [id]. -- If this variable is mapped to a Cminor local variable, the - corresponding Csharpminor memory block [b] must no longer exist in - Cminor ([f b = None]). Moreover, the content of block [b] must - match the value of [id] found in the Cminor local environment [te]. -- If this variable is mapped to a sub-block of the Cminor stack data - at offset [ofs], the address of this variable in Csharpminor [Vptr b - Int.zero] must match the address of the sub-block [Vptr sp (Int.repr - ofs)]. -*) - -Inductive match_var (f: meminj) (id: ident) - (e: Csharpminor.env) (m: mem) (te: env) (sp: block) : - var_info -> Prop := - | match_var_local: - forall chunk b v v', - PTree.get id e = Some (b, Vscalar chunk) -> - Mem.load chunk m b 0 = Some v -> - f b = None -> - PTree.get id te = Some v' -> - val_inject f v v' -> - match_var f id e m te sp (Var_local chunk) - | match_var_stack_scalar: - forall chunk ofs b, - PTree.get id e = Some (b, Vscalar chunk) -> - val_inject f (Vptr b Int.zero) (Vptr sp (Int.repr ofs)) -> - match_var f id e m te sp (Var_stack_scalar chunk ofs) - | match_var_stack_array: - forall ofs sz b, - PTree.get id e = Some (b, Varray sz) -> - val_inject f (Vptr b Int.zero) (Vptr sp (Int.repr ofs)) -> - match_var f id e m te sp (Var_stack_array ofs) - | match_var_global_scalar: - forall chunk, - PTree.get id e = None -> - PTree.get id (global_var_env prog) = Some (Vscalar chunk) -> - match_var f id e m te sp (Var_global_scalar chunk) - | match_var_global_array: - PTree.get id e = None -> - match_var f id e m te sp (Var_global_array). - -(** Matching between a Csharpminor environment [e] and a Cminor - environment [te]. The [lo] and [hi] parameters delimit the range - of addresses for the blocks referenced from [te]. *) - -Record match_env (f: meminj) (cenv: compilenv) - (e: Csharpminor.env) (m: mem) (te: env) (sp: block) - (lo hi: Z) : Prop := - mk_match_env { - -(** Each variable mentioned in the compilation environment must match - as defined above. *) - me_vars: - forall id, match_var f id e m te sp (PMap.get id cenv); - -(** The range [lo, hi] must not be empty. *) - me_low_high: - lo <= hi; - -(** Every block appearing in the Csharpminor environment [e] must be - in the range [lo, hi]. *) - me_bounded: - forall id b lv, PTree.get id e = Some(b, lv) -> lo <= b < hi; - -(** Distinct Csharpminor local variables must be mapped to distinct blocks. *) - me_inj: - forall id1 b1 lv1 id2 b2 lv2, - PTree.get id1 e = Some(b1, lv1) -> - PTree.get id2 e = Some(b2, lv2) -> - id1 <> id2 -> b1 <> b2; - -(** All blocks mapped to sub-blocks of the Cminor stack data must be in - the range [lo, hi]. *) - me_inv: - forall b delta, - f b = Some(sp, delta) -> lo <= b < hi; - -(** All Csharpminor blocks below [lo] (i.e. allocated before the blocks - referenced from [e]) must map to blocks that are below [sp] - (i.e. allocated before the stack data for the current Cminor function). *) - me_incr: - forall b tb delta, - f b = Some(tb, delta) -> b < lo -> tb < sp - }. - -(** Global environments match if the memory injection [f] leaves unchanged - the references to global symbols and functions. *) - -Record match_globalenvs (f: meminj) : Prop := - mk_match_globalenvs { - mg_symbols: - forall id b, - Genv.find_symbol ge id = Some b -> - f b = Some (b, 0) /\ Genv.find_symbol tge id = Some b; - mg_functions: - forall b, b < 0 -> f b = Some(b, 0) - }. - -(** Call stacks represent abstractly the execution state of the current - Csharpminor and Cminor functions, as well as the states of the - calling functions. A call stack is a list of frames, each frame - collecting information on the current execution state of a Csharpminor - function and its Cminor translation. *) - -Record frame : Set := - mkframe { - fr_cenv: compilenv; - fr_e: Csharpminor.env; - fr_te: env; - fr_sp: block; - fr_low: Z; - fr_high: Z - }. - -Definition callstack : Set := list frame. - -(** Matching of call stacks imply matching of environments for each of - the frames, plus matching of the global environments, plus disjointness - conditions over the memory blocks allocated for local variables - and Cminor stack data. *) - -Inductive match_callstack: meminj -> callstack -> Z -> Z -> mem -> Prop := - | mcs_nil: - forall f bound tbound m, - match_globalenvs f -> - match_callstack f nil bound tbound m - | mcs_cons: - forall f cenv e te sp lo hi cs bound tbound m, - hi <= bound -> - sp < tbound -> - match_env f cenv e m te sp lo hi -> - match_callstack f cs lo sp m -> - match_callstack f (mkframe cenv e te sp lo hi :: cs) bound tbound m. - -(** The remainder of this section is devoted to showing preservation - of the [match_callstack] invariant under various assignments and memory - stores. First: preservation by memory stores to ``mapped'' blocks - (block that have a counterpart in the Cminor execution). *) - -Lemma match_env_store_mapped: - forall f cenv e m1 m2 te sp lo hi chunk b ofs v, - f b <> None -> - store chunk m1 b ofs v = Some m2 -> - match_env f cenv e m1 te sp lo hi -> - match_env f cenv e m2 te sp lo hi. -Proof. - intros. inversion H1. constructor; auto. - (* vars *) - intros. generalize (me_vars0 id); intro. - inversion H2; econstructor; eauto. - rewrite <- H5. eapply load_store_other; eauto. - left. congruence. -Qed. - -Lemma match_callstack_mapped: - forall f cs bound tbound m1, - match_callstack f cs bound tbound m1 -> - forall chunk b ofs v m2, - f b <> None -> - store chunk m1 b ofs v = Some m2 -> - match_callstack f cs bound tbound m2. -Proof. - induction 1; intros; econstructor; eauto. - eapply match_env_store_mapped; eauto. -Qed. - -(** Preservation by assignment to a Csharpminor variable that is - translated to a Cminor local variable. The value being assigned - must be normalized with respect to the memory chunk of the variable, - in the following sense. *) - -(* -Definition val_normalized (chunk: memory_chunk) (v: val) : Prop := - exists v0, v = Val.load_result chunk v0. - -Lemma load_result_idem: - forall chunk v, - Val.load_result chunk (Val.load_result chunk v) = - Val.load_result chunk v. -Proof. - destruct chunk; destruct v; simpl; auto. - rewrite Int.cast8_signed_idem; auto. - rewrite Int.cast8_unsigned_idem; auto. - rewrite Int.cast16_signed_idem; auto. - rewrite Int.cast16_unsigned_idem; auto. - rewrite Float.singleoffloat_idem; auto. -Qed. - -Lemma load_result_normalized: - forall chunk v, - val_normalized chunk v -> Val.load_result chunk v = v. -Proof. - intros chunk v [v0 EQ]. rewrite EQ. apply load_result_idem. -Qed. -*) -Lemma match_env_store_local: - forall f cenv e m1 m2 te sp lo hi id b chunk v tv, - e!id = Some(b, Vscalar chunk) -> - val_inject f (Val.load_result chunk v) tv -> - store chunk m1 b 0 v = Some m2 -> - match_env f cenv e m1 te sp lo hi -> - match_env f cenv e m2 (PTree.set id tv te) sp lo hi. -Proof. - intros. inversion H2. constructor; auto. - intros. generalize (me_vars0 id0); intro. - inversion H3; subst. - (* var_local *) - case (peq id id0); intro. - (* the stored variable *) - subst id0. - change Csharpminor.var_kind with var_kind in H4. - rewrite H in H5. injection H5; clear H5; intros; subst b0 chunk0. - econstructor. eauto. - eapply load_store_same; eauto. auto. - rewrite PTree.gss. reflexivity. - auto. - (* a different variable *) - econstructor; eauto. - rewrite <- H6. eapply load_store_other; eauto. - rewrite PTree.gso; auto. - (* var_stack_scalar *) - econstructor; eauto. - (* var_stack_array *) - econstructor; eauto. - (* var_global_scalar *) - econstructor; eauto. - (* var_global_array *) - econstructor; eauto. -Qed. - -Lemma match_env_store_above: - forall f cenv e m1 m2 te sp lo hi chunk b v, - store chunk m1 b 0 v = Some m2 -> - hi <= b -> - match_env f cenv e m1 te sp lo hi -> - match_env f cenv e m2 te sp lo hi. -Proof. - intros. inversion H1; constructor; auto. - intros. generalize (me_vars0 id); intro. - inversion H2; econstructor; eauto. - rewrite <- H5. eapply load_store_other; eauto. - left. generalize (me_bounded0 _ _ _ H4). unfold block in *. omega. -Qed. - -Lemma match_callstack_store_above: - forall f cs bound tbound m1, - match_callstack f cs bound tbound m1 -> - forall chunk b v m2, - store chunk m1 b 0 v = Some m2 -> - bound <= b -> - match_callstack f cs bound tbound m2. -Proof. - induction 1; intros; econstructor; eauto. - eapply match_env_store_above with (b := b); eauto. omega. - eapply IHmatch_callstack; eauto. - inversion H1. omega. -Qed. - -Lemma match_callstack_store_local: - forall f cenv e te sp lo hi cs bound tbound m1 m2 id b chunk v tv, - e!id = Some(b, Vscalar chunk) -> - val_inject f (Val.load_result chunk v) tv -> - store chunk m1 b 0 v = Some m2 -> - match_callstack f (mkframe cenv e te sp lo hi :: cs) bound tbound m1 -> - match_callstack f (mkframe cenv e (PTree.set id tv te) sp lo hi :: cs) bound tbound m2. -Proof. - intros. inversion H2. constructor; auto. - eapply match_env_store_local; eauto. - eapply match_callstack_store_above; eauto. - inversion H16. - generalize (me_bounded0 _ _ _ H). omega. -Qed. - -(** A variant of [match_callstack_store_local] where the Cminor environment - [te] already associates to [id] a value that matches the assigned value. - In this case, [match_callstack] is preserved even if no assignment - takes place on the Cminor side. *) - -Lemma match_env_extensional: - forall f cenv e m te1 sp lo hi, - match_env f cenv e m te1 sp lo hi -> - forall te2, - (forall id, te2!id = te1!id) -> - match_env f cenv e m te2 sp lo hi. -Proof. - induction 1; intros; econstructor; eauto. - intros. generalize (me_vars0 id); intro. - inversion H0; econstructor; eauto. - rewrite H. auto. -Qed. - -Lemma match_callstack_store_local_unchanged: - forall f cenv e te sp lo hi cs bound tbound m1 m2 id b chunk v tv, - e!id = Some(b, Vscalar chunk) -> - val_inject f (Val.load_result chunk v) tv -> - store chunk m1 b 0 v = Some m2 -> - te!id = Some tv -> - match_callstack f (mkframe cenv e te sp lo hi :: cs) bound tbound m1 -> - match_callstack f (mkframe cenv e te sp lo hi :: cs) bound tbound m2. -Proof. - intros. inversion H3. constructor; auto. - apply match_env_extensional with (PTree.set id tv te). - eapply match_env_store_local; eauto. - intros. rewrite PTree.gsspec. - case (peq id0 id); intros. congruence. auto. - eapply match_callstack_store_above; eauto. - inversion H17. - generalize (me_bounded0 _ _ _ H). omega. -Qed. - -(** Preservation of [match_callstack] by freeing all blocks allocated - for local variables at function entry (on the Csharpminor side). *) - -Lemma match_callstack_incr_bound: - forall f cs bound tbound m, - match_callstack f cs bound tbound m -> - forall bound' tbound', - bound <= bound' -> tbound <= tbound' -> - match_callstack f cs bound' tbound' m. -Proof. - intros. inversion H; constructor; auto. omega. omega. -Qed. - -Lemma load_freelist: - forall fbl chunk m b ofs, - (forall b', In b' fbl -> b' <> b) -> - load chunk (free_list m fbl) b ofs = load chunk m b ofs. -Proof. - induction fbl; simpl; intros. - auto. - rewrite load_free. apply IHfbl. - intros. apply H. tauto. - apply sym_not_equal. apply H. tauto. -Qed. - -Lemma match_env_freelist: - forall f cenv e m te sp lo hi fbl, - match_env f cenv e m te sp lo hi -> - (forall b, In b fbl -> hi <= b) -> - match_env f cenv e (free_list m fbl) te sp lo hi. -Proof. - intros. inversion H. econstructor; eauto. - intros. generalize (me_vars0 id); intro. - inversion H1; econstructor; eauto. - rewrite <- H4. apply load_freelist. - intros. generalize (H0 _ H8); intro. - generalize (me_bounded0 _ _ _ H3). unfold block; omega. -Qed. - -Lemma match_callstack_freelist_rec: - forall f cs bound tbound m, - match_callstack f cs bound tbound m -> - forall fbl, - (forall b, In b fbl -> bound <= b) -> - match_callstack f cs bound tbound (free_list m fbl). -Proof. - induction 1; intros; constructor; auto. - eapply match_env_freelist; eauto. - intros. generalize (H3 _ H4). omega. - apply IHmatch_callstack. intros. - generalize (H3 _ H4). inversion H1. omega. -Qed. - -Lemma match_callstack_freelist: - forall f cenv e te sp lo hi cs bound tbound m fbl, - (forall b, In b fbl -> lo <= b) -> - match_callstack f (mkframe cenv e te sp lo hi :: cs) bound tbound m -> - match_callstack f cs bound tbound (free_list m fbl). -Proof. - intros. inversion H0. inversion H14. - apply match_callstack_incr_bound with lo sp. - apply match_callstack_freelist_rec. auto. - assumption. - omega. omega. -Qed. - -(** Preservation of [match_callstack] when allocating a block for - a local variable on the Csharpminor side. *) - -Lemma load_from_alloc_is_undef: - forall m1 chunk m2 b, - alloc m1 0 (size_chunk chunk) = (m2, b) -> - load chunk m2 b 0 = Some Vundef. -Proof. - intros. - assert (valid_block m2 b). eapply valid_new_block; eauto. - assert (low_bound m2 b <= 0). - generalize (low_bound_alloc _ _ b _ _ _ H). rewrite zeq_true. omega. - assert (0 + size_chunk chunk <= high_bound m2 b). - generalize (high_bound_alloc _ _ b _ _ _ H). rewrite zeq_true. omega. - elim (load_in_bounds _ _ _ _ H0 H1 H2). intros v LOAD. - assert (v = Vundef). eapply load_alloc_same; eauto. - congruence. -Qed. - -Lemma match_env_alloc_same: - forall m1 lv m2 b info f1 cenv1 e1 te sp lo id data tv, - alloc m1 0 (sizeof lv) = (m2, b) -> - match info with - | Var_local chunk => data = None /\ lv = Vscalar chunk - | Var_stack_scalar chunk pos => data = Some(sp, pos) /\ lv = Vscalar chunk - | Var_stack_array pos => data = Some(sp, pos) /\ exists sz, lv = Varray sz - | Var_global_scalar chunk => False - | Var_global_array => False - end -> - match_env f1 cenv1 e1 m1 te sp lo m1.(nextblock) -> - te!id = Some tv -> - let f2 := extend_inject b data f1 in - let cenv2 := PMap.set id info cenv1 in - let e2 := PTree.set id (b, lv) e1 in - inject_incr f1 f2 -> - match_env f2 cenv2 e2 m2 te sp lo m2.(nextblock). -Proof. - intros. - assert (b = m1.(nextblock)). - injection H; intros. auto. - assert (m2.(nextblock) = Zsucc m1.(nextblock)). - injection H; intros. rewrite <- H6; reflexivity. - inversion H1. constructor. - (* me_vars *) - intros id0. unfold cenv2. rewrite PMap.gsspec. case (peq id0 id); intros. - (* same var *) - subst id0. destruct info. - (* info = Var_local chunk *) - elim H0; intros. - apply match_var_local with b Vundef tv. - unfold e2; rewrite PTree.gss. congruence. - eapply load_from_alloc_is_undef; eauto. - rewrite H7 in H. unfold sizeof in H. eauto. - unfold f2, extend_inject, eq_block. rewrite zeq_true. auto. - auto. - constructor. - (* info = Var_stack_scalar chunk ofs *) - elim H0; intros. - apply match_var_stack_scalar with b. - unfold e2; rewrite PTree.gss. congruence. - eapply val_inject_ptr. - unfold f2, extend_inject, eq_block. rewrite zeq_true. eauto. - rewrite Int.add_commut. rewrite Int.add_zero. auto. - (* info = Var_stack_array z *) - elim H0; intros A [sz B]. - apply match_var_stack_array with sz b. - unfold e2; rewrite PTree.gss. congruence. - eapply val_inject_ptr. - unfold f2, extend_inject, eq_block. rewrite zeq_true. eauto. - rewrite Int.add_commut. rewrite Int.add_zero. auto. - (* info = Var_global *) - contradiction. - contradiction. - (* other vars *) - generalize (me_vars0 id0); intros. - inversion H6; econstructor; eauto. - unfold e2; rewrite PTree.gso; auto. - unfold f2, extend_inject, eq_block; rewrite zeq_false; auto. - generalize (me_bounded0 _ _ _ H8). unfold block in *; omega. - unfold e2; rewrite PTree.gso; eauto. - unfold e2; rewrite PTree.gso; eauto. - unfold e2; rewrite PTree.gso; eauto. - unfold e2; rewrite PTree.gso; eauto. - (* lo <= hi *) - unfold block in *; omega. - (* me_bounded *) - intros until lv0. unfold e2; rewrite PTree.gsspec. - case (peq id0 id); intros. - subst id0. inversion H6. subst b0. unfold block in *; omega. - generalize (me_bounded0 _ _ _ H6). rewrite H5. omega. - (* me_inj *) - intros until lv2. unfold e2; repeat rewrite PTree.gsspec. - case (peq id1 id); case (peq id2 id); intros. - congruence. - inversion H6. subst b1. rewrite H4. - generalize (me_bounded0 _ _ _ H7). unfold block; omega. - inversion H7. subst b2. rewrite H4. - generalize (me_bounded0 _ _ _ H6). unfold block; omega. - eauto. - (* me_inv *) - intros until delta. unfold f2, extend_inject, eq_block. - case (zeq b0 b); intros. - subst b0. rewrite H4; rewrite H5. omega. - generalize (me_inv0 _ _ H6). rewrite H5. omega. - (* me_incr *) - intros until delta. unfold f2, extend_inject, eq_block. - case (zeq b0 b); intros. - subst b0. unfold block in *; omegaContradiction. - eauto. -Qed. - -Lemma match_env_alloc_other: - forall f1 cenv e m1 m2 te sp lo hi chunk b data, - alloc m1 0 (sizeof chunk) = (m2, b) -> - match data with None => True | Some (b', delta') => sp < b' end -> - hi <= m1.(nextblock) -> - match_env f1 cenv e m1 te sp lo hi -> - let f2 := extend_inject b data f1 in - inject_incr f1 f2 -> - match_env f2 cenv e m2 te sp lo hi. -Proof. - intros. - assert (b = m1.(nextblock)). injection H; auto. - rewrite <- H4 in H1. - inversion H2. constructor; auto. - (* me_vars *) - intros. generalize (me_vars0 id); intro. - inversion H5; econstructor; eauto. - unfold f2, extend_inject, eq_block. rewrite zeq_false. auto. - generalize (me_bounded0 _ _ _ H7). unfold block in *; omega. - (* me_bounded *) - intros until delta. unfold f2, extend_inject, eq_block. - case (zeq b0 b); intros. rewrite H5 in H0. omegaContradiction. - eauto. - (* me_incr *) - intros until delta. unfold f2, extend_inject, eq_block. - case (zeq b0 b); intros. subst b0. omegaContradiction. - eauto. -Qed. - -Lemma match_callstack_alloc_other: - forall f1 cs bound tbound m1, - match_callstack f1 cs bound tbound m1 -> - forall lv m2 b data, - alloc m1 0 (sizeof lv) = (m2, b) -> - match data with None => True | Some (b', delta') => tbound <= b' end -> - bound <= m1.(nextblock) -> - let f2 := extend_inject b data f1 in - inject_incr f1 f2 -> - match_callstack f2 cs bound tbound m2. -Proof. - induction 1; intros. - constructor. - inversion H. constructor. - intros. auto. - intros. elim (mg_symbols0 _ _ H4); intros. - split; auto. elim (H3 b0); intros; congruence. - intros. generalize (mg_functions0 _ H4). elim (H3 b0); congruence. - constructor. auto. auto. - unfold f2; eapply match_env_alloc_other; eauto. - destruct data; auto. destruct p. omega. omega. - unfold f2; eapply IHmatch_callstack; eauto. - destruct data; auto. destruct p. omega. - inversion H1; omega. -Qed. - -Lemma match_callstack_alloc_left: - forall m1 lv m2 b info f1 cenv1 e1 te sp lo id data cs tv tbound, - alloc m1 0 (sizeof lv) = (m2, b) -> - match info with - | Var_local chunk => data = None /\ lv = Vscalar chunk - | Var_stack_scalar chunk pos => data = Some(sp, pos) /\ lv = Vscalar chunk - | Var_stack_array pos => data = Some(sp, pos) /\ exists sz, lv = Varray sz - | Var_global_scalar chunk => False - | Var_global_array => False - end -> - match_callstack f1 (mkframe cenv1 e1 te sp lo m1.(nextblock) :: cs) m1.(nextblock) tbound m1 -> - te!id = Some tv -> - let f2 := extend_inject b data f1 in - let cenv2 := PMap.set id info cenv1 in - let e2 := PTree.set id (b, lv) e1 in - inject_incr f1 f2 -> - match_callstack f2 (mkframe cenv2 e2 te sp lo m2.(nextblock) :: cs) m2.(nextblock) tbound m2. -Proof. - intros. inversion H1. constructor. omega. auto. - unfold f2, cenv2, e2. eapply match_env_alloc_same; eauto. - unfold f2; eapply match_callstack_alloc_other; eauto. - destruct info. - elim H0; intros. rewrite H19. auto. - elim H0; intros. rewrite H19. omega. - elim H0; intros. rewrite H19. omega. - contradiction. - contradiction. - inversion H17; omega. -Qed. - -Lemma match_callstack_alloc_right: - forall f cs bound tm1 m tm2 lo hi b, - alloc tm1 lo hi = (tm2, b) -> - match_callstack f cs bound tm1.(nextblock) m -> - match_callstack f cs bound tm2.(nextblock) m. -Proof. - intros. eapply match_callstack_incr_bound; eauto. omega. - injection H; intros. rewrite <- H2; simpl. omega. -Qed. - -Lemma match_env_alloc: - forall m1 l h m2 b tm1 tm2 tb f1 ce e te sp lo hi, - alloc m1 l h = (m2, b) -> - alloc tm1 l h = (tm2, tb) -> - match_env f1 ce e m1 te sp lo hi -> - hi <= m1.(nextblock) -> - sp < tm1.(nextblock) -> - let f2 := extend_inject b (Some(tb, 0)) f1 in - inject_incr f1 f2 -> - match_env f2 ce e m2 te sp lo hi. -Proof. - intros. - assert (BEQ: b = m1.(nextblock)). injection H; auto. - assert (TBEQ: tb = tm1.(nextblock)). injection H0; auto. - inversion H1. constructor; auto. - (* me_vars *) - intros. generalize (me_vars0 id); intro. inversion H5. - (* var_local *) - econstructor; eauto. - generalize (me_bounded0 _ _ _ H7). intro. - unfold f2, extend_inject. case (eq_block b0 b); intro. - subst b0. rewrite BEQ in H12. omegaContradiction. - auto. - (* var_stack_scalar *) - econstructor; eauto. - (* var_stack_array *) - econstructor; eauto. - (* var_global_scalar *) - econstructor; eauto. - (* var_global_array *) - econstructor; eauto. - (* me_bounded *) - intros until delta. unfold f2, extend_inject. case (eq_block b0 b); intro. - intro. injection H5; clear H5; intros. - rewrite H6 in TBEQ. rewrite TBEQ in H3. omegaContradiction. - eauto. - (* me_inj *) - intros until delta. unfold f2, extend_inject. case (eq_block b0 b); intros. - injection H5; clear H5; intros; subst b0 tb0 delta. - rewrite BEQ in H6. omegaContradiction. - eauto. -Qed. - -Lemma match_callstack_alloc_rec: - forall f1 cs bound tbound m1, - match_callstack f1 cs bound tbound m1 -> - forall l h m2 b tm1 tm2 tb, - alloc m1 l h = (m2, b) -> - alloc tm1 l h = (tm2, tb) -> - bound <= m1.(nextblock) -> - tbound <= tm1.(nextblock) -> - let f2 := extend_inject b (Some(tb, 0)) f1 in - inject_incr f1 f2 -> - match_callstack f2 cs bound tbound m2. -Proof. - induction 1; intros. - constructor. - inversion H. constructor. - intros. elim (mg_symbols0 _ _ H5); intros. - split; auto. elim (H4 b0); intros; congruence. - intros. generalize (mg_functions0 _ H5). elim (H4 b0); congruence. - constructor. auto. auto. - unfold f2. eapply match_env_alloc; eauto. omega. omega. - unfold f2; eapply IHmatch_callstack; eauto. - inversion H1; omega. - omega. -Qed. - -Lemma match_callstack_alloc: - forall f1 cs m1 tm1 l h m2 b tm2 tb, - match_callstack f1 cs m1.(nextblock) tm1.(nextblock) m1 -> - alloc m1 l h = (m2, b) -> - alloc tm1 l h = (tm2, tb) -> - let f2 := extend_inject b (Some(tb, 0)) f1 in - inject_incr f1 f2 -> - match_callstack f2 cs m2.(nextblock) tm2.(nextblock) m2. -Proof. - intros. unfold f2 in *. - apply match_callstack_incr_bound with m1.(nextblock) tm1.(nextblock). - eapply match_callstack_alloc_rec; eauto. omega. omega. - injection H0; intros; subst m2; simpl; omega. - injection H1; intros; subst tm2; simpl; omega. -Qed. - -(** [match_callstack] implies [match_globalenvs]. *) - -Lemma match_callstack_match_globalenvs: - forall f cs bound tbound m, - match_callstack f cs bound tbound m -> - match_globalenvs f. -Proof. - induction 1; eauto. -Qed. - -(** * Correctness of Cminor construction functions *) - -Hint Resolve eval_negint eval_negfloat eval_absfloat eval_intoffloat - eval_floatofint eval_floatofintu eval_notint eval_notbool - eval_cast8signed eval_cast8unsigned eval_cast16signed - eval_cast16unsigned eval_singleoffloat eval_add eval_add_ptr - eval_add_ptr_2 eval_sub eval_sub_ptr_int eval_sub_ptr_ptr - eval_mul eval_divs eval_mods eval_divu eval_modu - eval_and eval_or eval_xor eval_shl eval_shr eval_shru - eval_addf eval_subf eval_mulf eval_divf - eval_cmp eval_cmp_null_r eval_cmp_null_l eval_cmp_ptr - eval_cmpu eval_cmpf: evalexpr. - -Remark val_inject_val_of_bool: - forall f b, val_inject f (Val.of_bool b) (Val.of_bool b). -Proof. - intros; destruct b; unfold Val.of_bool, Vtrue, Vfalse; constructor. -Qed. - -Ltac TrivialOp := - match goal with - | [ |- exists x, _ /\ val_inject _ (Vint ?x) _ ] => - exists (Vint x); split; - [eauto with evalexpr | constructor] - | [ |- exists x, _ /\ val_inject _ (Vfloat ?x) _ ] => - exists (Vfloat x); split; - [eauto with evalexpr | constructor] - | [ |- exists x, _ /\ val_inject _ (Val.of_bool ?x) _ ] => - exists (Val.of_bool x); split; - [eauto with evalexpr | apply val_inject_val_of_bool] - | _ => idtac - end. - -Remark eval_compare_null_inv: - forall c i v, - Csharpminor.eval_compare_null c i = Some v -> - i = Int.zero /\ (c = Ceq /\ v = Vfalse \/ c = Cne /\ v = Vtrue). -Proof. - intros until v. unfold Csharpminor.eval_compare_null. - predSpec Int.eq Int.eq_spec i Int.zero. - case c; intro EQ; simplify_eq EQ; intro; subst v; tauto. - congruence. -Qed. - -(** Correctness of [make_op]. The generated Cminor code evaluates - to a value that matches the result value of the Csharpminor operation, - provided arguments match pairwise ([val_list_inject f] hypothesis). *) - -Lemma make_op_correct: - forall al a op vl m2 v sp le te1 tm1 t te2 tm2 tvl f, - make_op op al = Some a -> - Csharpminor.eval_operation op vl m2 = Some v -> - eval_exprlist tge (Vptr sp Int.zero) le te1 tm1 al t te2 tm2 tvl -> - val_list_inject f vl tvl -> - mem_inject f m2 tm2 -> - exists tv, - eval_expr tge (Vptr sp Int.zero) le te1 tm1 a t te2 tm2 tv - /\ val_inject f v tv. -Proof. - intros. - destruct al as [ | a1 al]; - [idtac | destruct al as [ | a2 al]; - [idtac | destruct al as [ | a3 al]]]; - simpl in H; try discriminate. - (* Constant operators *) - inversion H1. subst sp0 le0 e m te2 tm1 tvl. - inversion H2. subst vl. - destruct op; simplify_eq H; intro; subst a; - simpl in H0; injection H0; intro; subst v. - (* intconst *) - TrivialOp. econstructor. constructor. reflexivity. - (* floatconst *) - TrivialOp. econstructor. constructor. reflexivity. - (* Unary operators *) - inversion H1; clear H1; subst. - inversion H11; clear H11; subst. - rewrite E0_right. - inversion H2; clear H2; subst. inversion H8; clear H8; subst. - destruct op; simplify_eq H; intro; subst a; - simpl in H0; destruct v1; simplify_eq H0; intro; subst v; - inversion H7; subst v0; - TrivialOp. - change (Vint (Int.cast8unsigned i)) with (Val.cast8unsigned (Vint i)). eauto with evalexpr. - change (Vint (Int.cast8signed i)) with (Val.cast8signed (Vint i)). eauto with evalexpr. - change (Vint (Int.cast16unsigned i)) with (Val.cast16unsigned (Vint i)). eauto with evalexpr. - change (Vint (Int.cast16signed i)) with (Val.cast16signed (Vint i)). eauto with evalexpr. - change (Vfloat (Float.singleoffloat f0)) with (Val.singleoffloat (Vfloat f0)). eauto with evalexpr. - (* Binary operations *) - inversion H1; clear H1; subst. inversion H11; clear H11; subst. - inversion H12; clear H12; subst. rewrite E0_right. - inversion H2; clear H2; subst. inversion H9; clear H9; subst. - inversion H10; clear H10; subst. - destruct op; simplify_eq H; intro; subst a; - simpl in H0; destruct v2; destruct v3; simplify_eq H0; intro; try subst v; - inversion H7; inversion H8; subst v0; subst v1; - TrivialOp. - (* add int ptr *) - exists (Vptr b2 (Int.add ofs2 i)); split. - eauto with evalexpr. apply val_inject_ptr with x. auto. - subst ofs2. repeat rewrite Int.add_assoc. decEq. apply Int.add_commut. - (* add ptr int *) - exists (Vptr b2 (Int.add ofs2 i0)); split. - eauto with evalexpr. apply val_inject_ptr with x. auto. - subst ofs2. repeat rewrite Int.add_assoc. decEq. apply Int.add_commut. - (* sub ptr int *) - exists (Vptr b2 (Int.sub ofs2 i0)); split. - eauto with evalexpr. apply val_inject_ptr with x. auto. - subst ofs2. apply Int.sub_add_l. - (* sub ptr ptr *) - destruct (eq_block b b0); simplify_eq H0; intro; subst v; subst b0. - assert (b4 = b2). congruence. subst b4. - exists (Vint (Int.sub ofs3 ofs2)); split. - eauto with evalexpr. - subst ofs2 ofs3. replace x0 with x. rewrite Int.sub_shifted. constructor. - congruence. - (* divs *) - generalize (Int.eq_spec i0 Int.zero); destruct (Int.eq i0 Int.zero); intro; - simplify_eq H0; intro; subst v. TrivialOp. - (* divu *) - generalize (Int.eq_spec i0 Int.zero); destruct (Int.eq i0 Int.zero); intro; - simplify_eq H0; intro; subst v. TrivialOp. - (* mods *) - generalize (Int.eq_spec i0 Int.zero); destruct (Int.eq i0 Int.zero); intro; - simplify_eq H0; intro; subst v. TrivialOp. - (* modu *) - generalize (Int.eq_spec i0 Int.zero); destruct (Int.eq i0 Int.zero); intro; - simplify_eq H0; intro; subst v. TrivialOp. - (* shl *) - caseEq (Int.ltu i0 (Int.repr 32)); intro EQ; rewrite EQ in H0; - simplify_eq H0; intro; subst v. TrivialOp. - (* shr *) - caseEq (Int.ltu i0 (Int.repr 32)); intro EQ; rewrite EQ in H0; - simplify_eq H0; intro; subst v. TrivialOp. - (* shru *) - caseEq (Int.ltu i0 (Int.repr 32)); intro EQ; rewrite EQ in H0; - simplify_eq H0; intro; subst v. TrivialOp. - (* cmp int ptr *) - elim (eval_compare_null_inv _ _ _ H0); intros; subst i1 i. - exists v; split. eauto with evalexpr. - elim H12; intros [A B]; subst v; unfold Vtrue, Vfalse; constructor. - (* cmp ptr int *) - elim (eval_compare_null_inv _ _ _ H0); intros; subst i1 i0. - exists v; split. eauto with evalexpr. - elim H12; intros [A B]; subst v; unfold Vtrue, Vfalse; constructor. - (* cmp ptr ptr *) - caseEq (valid_pointer m2 b (Int.signed i) && valid_pointer m2 b0 (Int.signed i0)); - intro EQ; rewrite EQ in H0; try discriminate. - destruct (eq_block b b0); simplify_eq H0; intro; subst v b0. - assert (b4 = b2); [congruence|subst b4]. - assert (x0 = x); [congruence|subst x0]. - elim (andb_prop _ _ EQ); intros. - exists (Val.of_bool (Int.cmp c ofs3 ofs2)); split. - eauto with evalexpr. - subst ofs2 ofs3. rewrite Int.translate_cmp. - apply val_inject_val_of_bool. - eapply valid_pointer_inject_no_overflow; eauto. - eapply valid_pointer_inject_no_overflow; eauto. -Qed. - -(** Correctness of [make_cast]. Note that the resulting Cminor value is - normalized according to the given memory chunk. *) - -Lemma make_cast_correct: - forall f sp le te1 tm1 a t te2 tm2 v chunk tv, - eval_expr tge (Vptr sp Int.zero) le te1 tm1 a t te2 tm2 tv -> - val_inject f v tv -> - exists tv', - eval_expr tge (Vptr sp Int.zero) le - te1 tm1 (make_cast chunk a) - t te2 tm2 tv' - /\ val_inject f (Val.load_result chunk v) tv'. -Proof. - intros. destruct chunk. - - exists (Val.cast8signed tv). - split. apply eval_cast8signed; auto. - inversion H0; simpl; constructor. - - exists (Val.cast8unsigned tv). - split. apply eval_cast8unsigned; auto. - inversion H0; simpl; constructor. - - exists (Val.cast16signed tv). - split. apply eval_cast16signed; auto. - inversion H0; simpl; constructor. - - exists (Val.cast16unsigned tv). - split. apply eval_cast16unsigned; auto. - inversion H0; simpl; constructor. - - exists tv. - split. simpl; auto. - inversion H0; simpl; econstructor; eauto. - - exists (Val.singleoffloat tv). - split. apply eval_singleoffloat; auto. - inversion H0; simpl; constructor. - - exists tv. - split. simpl; auto. - inversion H0; simpl; constructor. -Qed. - -Lemma make_stackaddr_correct: - forall sp le te tm ofs, - eval_expr tge (Vptr sp Int.zero) le - te tm (make_stackaddr ofs) - E0 te tm (Vptr sp (Int.repr ofs)). -Proof. - intros; unfold make_stackaddr. - eapply eval_Eop. econstructor. simpl. decEq. decEq. - rewrite Int.add_commut. apply Int.add_zero. -Qed. - -Lemma make_globaladdr_correct: - forall sp le te tm id b, - Genv.find_symbol tge id = Some b -> - eval_expr tge (Vptr sp Int.zero) le - te tm (make_globaladdr id) - E0 te tm (Vptr b Int.zero). -Proof. - intros; unfold make_globaladdr. - eapply eval_Eop. econstructor. simpl. rewrite H. auto. -Qed. - -(** Correctness of [make_load] and [make_store]. *) - -Lemma make_load_correct: - forall sp le te1 tm1 a t te2 tm2 va chunk v, - eval_expr tge (Vptr sp Int.zero) le te1 tm1 a t te2 tm2 va -> - Mem.loadv chunk tm2 va = Some v -> - eval_expr tge (Vptr sp Int.zero) le - te1 tm1 (make_load chunk a) - t te2 tm2 v. -Proof. - intros; unfold make_load. - eapply eval_load; eauto. -Qed. - -Lemma store_arg_content_inject: - forall f sp le te1 tm1 a t te2 tm2 v va chunk, - eval_expr tge (Vptr sp Int.zero) le te1 tm1 a t te2 tm2 va -> - val_inject f v va -> - exists vb, - eval_expr tge (Vptr sp Int.zero) le te1 tm1 (store_arg chunk a) t te2 tm2 vb - /\ val_content_inject f (mem_chunk chunk) v vb. -Proof. - intros. - assert (exists vb, - eval_expr tge (Vptr sp Int.zero) le te1 tm1 a t te2 tm2 vb - /\ val_content_inject f (mem_chunk chunk) v vb). - exists va; split. assumption. constructor. assumption. - inversion H; clear H; subst; simpl; trivial. - inversion H2; clear H2; subst; trivial. - inversion H4; clear H4; subst; trivial. - rewrite E0_right. rewrite E0_right in H1. - destruct op; trivial; destruct chunk; trivial; - exists v0; (split; [auto| - simpl in H3; inversion H3; clear H3; subst va; - destruct v0; simpl in H0; inversion H0; subst; try (constructor; constructor)]). - apply val_content_inject_8. apply Int.cast8_unsigned_signed. - apply val_content_inject_8. apply Int.cast8_unsigned_idem. - apply val_content_inject_16. apply Int.cast16_unsigned_signed. - apply val_content_inject_16. apply Int.cast16_unsigned_idem. - apply val_content_inject_32. apply Float.singleoffloat_idem. -Qed. - -Lemma make_store_correct: - forall f sp te1 tm1 addr te2 tm2 tvaddr rhs te3 tm3 tvrhs - chunk vrhs m3 vaddr m4 t1 t2, - eval_expr tge (Vptr sp Int.zero) nil - te1 tm1 addr t1 te2 tm2 tvaddr -> - eval_expr tge (Vptr sp Int.zero) nil - te2 tm2 rhs t2 te3 tm3 tvrhs -> - Mem.storev chunk m3 vaddr vrhs = Some m4 -> - mem_inject f m3 tm3 -> - val_inject f vaddr tvaddr -> - val_inject f vrhs tvrhs -> - exists tm4, - exec_stmt tge (Vptr sp Int.zero) - te1 tm1 (make_store chunk addr rhs) - (t1**t2) te3 tm4 Out_normal - /\ mem_inject f m4 tm4 - /\ nextblock tm4 = nextblock tm3. -Proof. - intros. unfold make_store. - exploit store_arg_content_inject. eexact H0. eauto. - intros [tv [EVAL VCINJ]]. - exploit storev_mapped_inject_1; eauto. - intros [tm4 [STORE MEMINJ]]. - exploit eval_store. eexact H. eexact EVAL. eauto. - intro EVALSTORE. - exists tm4. - split. apply exec_Sexpr with tv. auto. - split. auto. - unfold storev in STORE; destruct tvaddr; try discriminate. - exploit store_inv; eauto. simpl. tauto. -Qed. - -(** Correctness of the variable accessors [var_get], [var_set] - and [var_addr]. *) - -Lemma var_get_correct: - forall cenv id a f e te sp lo hi m cs tm b chunk v le, - var_get cenv id = Some a -> - match_callstack f (mkframe cenv e te sp lo hi :: cs) m.(nextblock) tm.(nextblock) m -> - mem_inject f m tm -> - eval_var_ref prog e id b chunk -> - load chunk m b 0 = Some v -> - exists tv, - eval_expr tge (Vptr sp Int.zero) le te tm a E0 te tm tv /\ - val_inject f v tv. -Proof. - unfold var_get; intros. - assert (match_var f id e m te sp cenv!!id). - inversion H0. inversion H17. auto. - inversion H4; subst; rewrite <- H5 in H; inversion H; subst. - (* var_local *) - inversion H2; [subst|congruence]. - exists v'; split. - apply eval_Evar. auto. - replace v with v0. auto. congruence. - (* var_stack_scalar *) - inversion H2; [subst|congruence]. - assert (b0 = b). congruence. subst b0. - assert (chunk0 = chunk). congruence. subst chunk0. - exploit loadv_inject; eauto. - unfold loadv. eexact H3. - intros [tv [LOAD INJ]]. - exists tv; split. - eapply make_load_correct; eauto. eapply make_stackaddr_correct; eauto. - auto. - (* var_global_scalar *) - inversion H2; [congruence|subst]. - assert (match_globalenvs f). eapply match_callstack_match_globalenvs; eauto. - inversion H11. destruct (mg_symbols0 _ _ H9) as [A B]. - assert (chunk0 = chunk). congruence. subst chunk0. - assert (loadv chunk m (Vptr b Int.zero) = Some v). assumption. - assert (val_inject f (Vptr b Int.zero) (Vptr b Int.zero)). - econstructor; eauto. - generalize (loadv_inject _ _ _ _ _ _ _ H1 H12 H13). - intros [tv [LOAD INJ]]. - exists tv; split. - eapply make_load_correct; eauto. eapply make_globaladdr_correct; eauto. - auto. -Qed. - -Lemma var_addr_correct: - forall cenv id a f e te sp lo hi m cs tm b le, - match_callstack f (mkframe cenv e te sp lo hi :: cs) m.(nextblock) tm.(nextblock) m -> - var_addr cenv id = Some a -> - eval_var_addr prog e id b -> - exists tv, - eval_expr tge (Vptr sp Int.zero) le te tm a E0 te tm tv /\ - val_inject f (Vptr b Int.zero) tv. -Proof. - unfold var_addr; intros. - assert (match_var f id e m te sp cenv!!id). - inversion H. inversion H15. auto. - inversion H2; subst; rewrite <- H3 in H0; inversion H0; subst; clear H0. - (* var_stack_scalar *) - inversion H1; [subst|congruence]. - exists (Vptr sp (Int.repr ofs)); split. - eapply make_stackaddr_correct. - replace b with b0. auto. congruence. - (* var_stack_array *) - inversion H1; [subst|congruence]. - exists (Vptr sp (Int.repr ofs)); split. - eapply make_stackaddr_correct. - replace b with b0. auto. congruence. - (* var_global_scalar *) - inversion H1; [congruence|subst]. - assert (match_globalenvs f). eapply match_callstack_match_globalenvs; eauto. - inversion H7. destruct (mg_symbols0 _ _ H6) as [A B]. - exists (Vptr b Int.zero); split. - eapply make_globaladdr_correct. eauto. - econstructor; eauto. - (* var_global_array *) - inversion H1; [congruence|subst]. - assert (match_globalenvs f). eapply match_callstack_match_globalenvs; eauto. - inversion H6. destruct (mg_symbols0 _ _ H5) as [A B]. - exists (Vptr b Int.zero); split. - eapply make_globaladdr_correct. eauto. - econstructor; eauto. -Qed. - -Lemma var_set_correct: - forall cenv id rhs a f e te2 sp lo hi m2 cs tm2 te1 tm1 tv b chunk v m3 t, - var_set cenv id rhs = Some a -> - match_callstack f (mkframe cenv e te2 sp lo hi :: cs) m2.(nextblock) tm2.(nextblock) m2 -> - eval_expr tge (Vptr sp Int.zero) nil te1 tm1 rhs t te2 tm2 tv -> - val_inject f v tv -> - mem_inject f m2 tm2 -> - eval_var_ref prog e id b chunk -> - store chunk m2 b 0 v = Some m3 -> - exists te3, exists tm3, - exec_stmt tge (Vptr sp Int.zero) te1 tm1 a t te3 tm3 Out_normal /\ - mem_inject f m3 tm3 /\ - match_callstack f (mkframe cenv e te3 sp lo hi :: cs) m3.(nextblock) tm3.(nextblock) m3. -Proof. - unfold var_set; intros. - assert (NEXTBLOCK: nextblock m3 = nextblock m2). - exploit store_inv; eauto. simpl; tauto. - inversion H0. subst f0 cenv0 e0 te sp0 lo0 hi0 cs0 bound tbound m. - assert (match_var f id e m2 te2 sp cenv!!id). inversion H19; auto. - inversion H6; subst; rewrite <- H7 in H; inversion H; subst; clear H. - (* var_local *) - inversion H4; [subst|congruence]. - assert (b0 = b). congruence. subst b0. - assert (chunk0 = chunk). congruence. subst chunk0. - exploit make_cast_correct; eauto. - intros [tv' [EVAL INJ]]. - exists (PTree.set id tv' te2); exists tm2. - split. eapply exec_Sexpr. eapply eval_Eassign. eauto. - split. eapply store_unmapped_inject; eauto. - rewrite NEXTBLOCK. eapply match_callstack_store_local; eauto. - (* var_stack_scalar *) - inversion H4; [subst|congruence]. - assert (b0 = b). congruence. subst b0. - assert (chunk0 = chunk). congruence. subst chunk0. - assert (storev chunk m2 (Vptr b Int.zero) v = Some m3). assumption. - exploit make_store_correct. - eapply make_stackaddr_correct. - eauto. eauto. eauto. eauto. eauto. - rewrite E0_left. intros [tm3 [EVAL [MEMINJ TNEXTBLOCK]]]. - exists te2; exists tm3. - split. auto. split. auto. - rewrite NEXTBLOCK; rewrite TNEXTBLOCK. - eapply match_callstack_mapped; eauto. - inversion H9; congruence. - (* var_global_scalar *) - inversion H4; [congruence|subst]. - assert (chunk0 = chunk). congruence. subst chunk0. - assert (storev chunk m2 (Vptr b Int.zero) v = Some m3). assumption. - assert (match_globalenvs f). eapply match_callstack_match_globalenvs; eauto. - inversion H13. destruct (mg_symbols0 _ _ H10) as [A B]. - exploit make_store_correct. - eapply make_globaladdr_correct; eauto. - eauto. eauto. eauto. eauto. eauto. - rewrite E0_left. intros [tm3 [EVAL [MEMINJ TNEXTBLOCK]]]. - exists te2; exists tm3. - split. auto. split. auto. - rewrite NEXTBLOCK; rewrite TNEXTBLOCK. - eapply match_callstack_mapped; eauto. congruence. -Qed. - -(** * Correctness of stack allocation of local variables *) - -(** This section shows the correctness of the translation of Csharpminor - local variables, either as Cminor local variables or as sub-blocks - of the Cminor stack data. This is the most difficult part of the proof. *) - -Remark assign_variables_incr: - forall atk vars cenv sz cenv' sz', - assign_variables atk vars (cenv, sz) = (cenv', sz') -> sz <= sz'. -Proof. - induction vars; intros until sz'; simpl. - intro. replace sz' with sz. omega. congruence. - destruct a. destruct v. case (Identset.mem i atk); intros. - generalize (IHvars _ _ _ _ H). - generalize (size_chunk_pos m). intro. - generalize (align_le sz (size_chunk m) H0). omega. - eauto. - intro. generalize (IHvars _ _ _ _ H). - assert (8 > 0). omega. generalize (align_le sz 8 H0). - assert (0 <= Zmax 0 z). apply Zmax_bound_l. omega. - omega. -Qed. - -Lemma match_callstack_alloc_variables_rec: - forall tm sp cenv' sz' te lo cs atk, - valid_block tm sp -> - low_bound tm sp = 0 -> - high_bound tm sp = sz' -> - sz' <= Int.max_signed -> - forall e m vars e' m' lb, - alloc_variables e m vars e' m' lb -> - forall f cenv sz, - assign_variables atk vars (cenv, sz) = (cenv', sz') -> - match_callstack f (mkframe cenv e te sp lo m.(nextblock) :: cs) - m.(nextblock) tm.(nextblock) m -> - mem_inject f m tm -> - 0 <= sz -> - (forall b delta, f b = Some(sp, delta) -> high_bound m b + delta <= sz) -> - (forall id lv, In (id, lv) vars -> te!id <> None) -> - exists f', - inject_incr f f' - /\ mem_inject f' m' tm - /\ match_callstack f' (mkframe cenv' e' te sp lo m'.(nextblock) :: cs) - m'.(nextblock) tm.(nextblock) m'. -Proof. - intros until atk. intros VB LB HB NOOV. - induction 1. - (* base case *) - intros. simpl in H. inversion H; subst cenv sz. - exists f. split. apply inject_incr_refl. split. auto. auto. - (* inductive case *) - intros until sz. - change (assign_variables atk ((id, lv) :: vars) (cenv, sz)) - with (assign_variables atk vars (assign_variable atk (id, lv) (cenv, sz))). - caseEq (assign_variable atk (id, lv) (cenv, sz)). - intros cenv1 sz1 ASV1 ASVS MATCH MINJ SZPOS BOUND DEFINED. - assert (DEFINED1: forall id0 lv0, In (id0, lv0) vars -> te!id0 <> None). - intros. eapply DEFINED. simpl. right. eauto. - assert (exists tv, te!id = Some tv). - assert (te!id <> None). eapply DEFINED. simpl; left; auto. - destruct (te!id). exists v; auto. congruence. - elim H1; intros tv TEID; clear H1. - generalize ASV1. unfold assign_variable. - caseEq lv. - (* 1. lv = LVscalar chunk *) - intros chunk LV. case (Identset.mem id atk). - (* 1.1 info = Var_stack_scalar chunk ... *) - set (ofs := align sz (size_chunk chunk)). - intro EQ; injection EQ; intros; clear EQ. - set (f1 := extend_inject b1 (Some (sp, ofs)) f). - generalize (size_chunk_pos chunk); intro SIZEPOS. - generalize (align_le sz (size_chunk chunk) SIZEPOS). fold ofs. intro SZOFS. - assert (mem_inject f1 m1 tm /\ inject_incr f f1). - assert (Int.min_signed < 0). compute; auto. - generalize (assign_variables_incr _ _ _ _ _ _ ASVS). intro. - unfold f1; eapply alloc_mapped_inject; eauto. - omega. omega. omega. omega. unfold sizeof; rewrite LV. omega. - intros. left. generalize (BOUND _ _ H5). omega. - elim H3; intros MINJ1 INCR1; clear H3. - exploit IHalloc_variables; eauto. - unfold f1; rewrite <- H2; eapply match_callstack_alloc_left; eauto. - rewrite <- H1. omega. - intros until delta; unfold f1, extend_inject, eq_block. - rewrite (high_bound_alloc _ _ b _ _ _ H). - case (zeq b b1); intros. - inversion H3. unfold sizeof; rewrite LV. omega. - generalize (BOUND _ _ H3). omega. - intros [f' [INCR2 [MINJ2 MATCH2]]]. - exists f'; intuition. eapply inject_incr_trans; eauto. - (* 1.2 info = Var_local chunk *) - intro EQ; injection EQ; intros; clear EQ. subst sz1. - exploit alloc_unmapped_inject; eauto. - set (f1 := extend_inject b1 None f). intros [MINJ1 INCR1]. - exploit IHalloc_variables; eauto. - unfold f1; rewrite <- H2; eapply match_callstack_alloc_left; eauto. - intros until delta; unfold f1, extend_inject, eq_block. - rewrite (high_bound_alloc _ _ b _ _ _ H). - case (zeq b b1); intros. discriminate. - eapply BOUND; eauto. - intros [f' [INCR2 [MINJ2 MATCH2]]]. - exists f'; intuition. eapply inject_incr_trans; eauto. - (* 2. lv = LVarray dim, info = Var_stack_array *) - intros dim LV EQ. injection EQ; clear EQ; intros. - assert (0 <= Zmax 0 dim). apply Zmax1. - assert (8 > 0). omega. - generalize (align_le sz 8 H4). intro. - set (ofs := align sz 8) in *. - set (f1 := extend_inject b1 (Some (sp, ofs)) f). - assert (mem_inject f1 m1 tm /\ inject_incr f f1). - assert (Int.min_signed < 0). compute; auto. - generalize (assign_variables_incr _ _ _ _ _ _ ASVS). intro. - unfold f1; eapply alloc_mapped_inject; eauto. - omega. omega. omega. omega. unfold sizeof; rewrite LV. omega. - intros. left. generalize (BOUND _ _ H8). omega. - destruct H6 as [MINJ1 INCR1]. - exploit IHalloc_variables; eauto. - unfold f1; rewrite <- H2; eapply match_callstack_alloc_left; eauto. - rewrite <- H1. omega. - intros until delta; unfold f1, extend_inject, eq_block. - rewrite (high_bound_alloc _ _ b _ _ _ H). - case (zeq b b1); intros. - inversion H6. unfold sizeof; rewrite LV. omega. - generalize (BOUND _ _ H6). omega. - intros [f' [INCR2 [MINJ2 MATCH2]]]. - exists f'; intuition. eapply inject_incr_trans; eauto. -Qed. - -Lemma set_params_defined: - forall params args id, - In id params -> (set_params args params)!id <> None. -Proof. - induction params; simpl; intros. - elim H. - destruct args. - rewrite PTree.gsspec. case (peq id a); intro. - congruence. eapply IHparams. elim H; intro. congruence. auto. - rewrite PTree.gsspec. case (peq id a); intro. - congruence. eapply IHparams. elim H; intro. congruence. auto. -Qed. - -Lemma set_locals_defined: - forall e vars id, - In id vars \/ e!id <> None -> (set_locals vars e)!id <> None. -Proof. - induction vars; simpl; intros. - tauto. - rewrite PTree.gsspec. case (peq id a); intro. - congruence. - apply IHvars. assert (a <> id). congruence. tauto. -Qed. - -Lemma set_locals_params_defined: - forall args params vars id, - In id (params ++ vars) -> - (set_locals vars (set_params args params))!id <> None. -Proof. - intros. apply set_locals_defined. - elim (in_app_or _ _ _ H); intro. - right. apply set_params_defined; auto. - left; auto. -Qed. - -(** Preservation of [match_callstack] by simultaneous allocation - of Csharpminor local variables and of the Cminor stack data block. *) - -Lemma match_callstack_alloc_variables: - forall fn cenv sz m e m' lb tm tm' sp f cs targs, - build_compilenv gce fn = (cenv, sz) -> - sz <= Int.max_signed -> - alloc_variables Csharpminor.empty_env m (fn_variables fn) e m' lb -> - Mem.alloc tm 0 sz = (tm', sp) -> - match_callstack f cs m.(nextblock) tm.(nextblock) m -> - mem_inject f m tm -> - let tparams := List.map (@fst ident memory_chunk) fn.(Csharpminor.fn_params) in - let tvars := List.map (@fst ident var_kind) fn.(Csharpminor.fn_vars) in - let te := set_locals tvars (set_params targs tparams) in - exists f', - inject_incr f f' - /\ mem_inject f' m' tm' - /\ match_callstack f' (mkframe cenv e te sp m.(nextblock) m'.(nextblock) :: cs) - m'.(nextblock) tm'.(nextblock) m'. -Proof. - intros. - assert (SP: sp = nextblock tm). injection H2; auto. - unfold build_compilenv in H. - eapply match_callstack_alloc_variables_rec with (sz' := sz); eauto. - eapply valid_new_block; eauto. - rewrite (low_bound_alloc _ _ sp _ _ _ H2). apply zeq_true. - rewrite (high_bound_alloc _ _ sp _ _ _ H2). apply zeq_true. - (* match_callstack *) - constructor. omega. change (valid_block tm' sp). eapply valid_new_block; eauto. - constructor. - (* me_vars *) - intros. generalize (global_compilenv_charact id). - destruct (gce!!id); intro; try contradiction. - constructor. - unfold Csharpminor.empty_env. apply PTree.gempty. auto. - constructor. - unfold Csharpminor.empty_env. apply PTree.gempty. - (* me_low_high *) - omega. - (* me_bounded *) - intros until lv. unfold Csharpminor.empty_env. rewrite PTree.gempty. congruence. - (* me_inj *) - intros until lv2. unfold Csharpminor.empty_env; rewrite PTree.gempty; congruence. - (* me_inv *) - intros. exploit mi_mappedblocks; eauto. intros [A B]. - elim (fresh_block_alloc _ _ _ _ _ H2 A). - (* me_incr *) - intros. exploit mi_mappedblocks; eauto. intros [A B]. - rewrite SP; auto. - rewrite SP; auto. - eapply alloc_right_inject; eauto. - omega. - intros. exploit mi_mappedblocks; eauto. intros [A B]. - unfold block in SP; omegaContradiction. - (* defined *) - intros. unfold te. apply set_locals_params_defined. - unfold tparams, tvars. unfold fn_variables in H5. - change Csharpminor.fn_params with Csharpminor.fn_params in H5. - change Csharpminor.fn_vars with Csharpminor.fn_vars in H5. - elim (in_app_or _ _ _ H5); intros. - elim (list_in_map_inv _ _ _ H6). intros x [A B]. - apply in_or_app; left. inversion A. apply List.in_map. auto. - apply in_or_app; right. - change id with (fst (id, lv)). apply List.in_map; auto. -Qed. - -(** Characterization of the range of addresses for the blocks allocated - to hold Csharpminor local variables. *) - -Lemma alloc_variables_nextblock_incr: - forall e1 m1 vars e2 m2 lb, - alloc_variables e1 m1 vars e2 m2 lb -> - nextblock m1 <= nextblock m2. -Proof. - induction 1; intros. - omega. - inversion H; subst m1; simpl in IHalloc_variables. omega. -Qed. - -Lemma alloc_variables_list_block: - forall e1 m1 vars e2 m2 lb, - alloc_variables e1 m1 vars e2 m2 lb -> - forall b, m1.(nextblock) <= b < m2.(nextblock) <-> In b lb. -Proof. - induction 1; intros. - simpl; split; intro. omega. contradiction. - elim (IHalloc_variables b); intros A B. - assert (nextblock m = b1). injection H; intros. auto. - assert (nextblock m1 = Zsucc (nextblock m)). - injection H; intros; subst m1; reflexivity. - simpl; split; intro. - assert (nextblock m = b \/ nextblock m1 <= b < nextblock m2). - unfold block; rewrite H2; omega. - elim H4; intro. left; congruence. right; auto. - elim H3; intro. subst b b1. - generalize (alloc_variables_nextblock_incr _ _ _ _ _ _ H0). - rewrite H2. omega. - generalize (B H4). rewrite H2. omega. -Qed. - -(** Correctness of the code generated by [store_parameters] - to store in memory the values of parameters that are stack-allocated. *) - -Inductive vars_vals_match: - meminj -> list (ident * memory_chunk) -> list val -> env -> Prop := - | vars_vals_nil: - forall f te, - vars_vals_match f nil nil te - | vars_vals_cons: - forall f te id chunk vars v vals tv, - te!id = Some tv -> - val_inject f v tv -> - vars_vals_match f vars vals te -> - vars_vals_match f ((id, chunk) :: vars) (v :: vals) te. - -Lemma vars_vals_match_extensional: - forall f vars vals te, - vars_vals_match f vars vals te -> - forall te', - (forall id lv, In (id, lv) vars -> te'!id = te!id) -> - vars_vals_match f vars vals te'. -Proof. - induction 1; intros. - constructor. - econstructor; eauto. rewrite <- H. eapply H2. left. reflexivity. - apply IHvars_vals_match. intros. eapply H2; eauto. right. eauto. -Qed. - -Lemma store_parameters_correct: - forall e m1 params vl m2, - bind_parameters e m1 params vl m2 -> - forall f te1 cenv sp lo hi cs tm1, - vars_vals_match f params vl te1 -> - list_norepet (List.map (@fst ident memory_chunk) params) -> - mem_inject f m1 tm1 -> - match_callstack f (mkframe cenv e te1 sp lo hi :: cs) m1.(nextblock) tm1.(nextblock) m1 -> - exists te2, exists tm2, - exec_stmt tge (Vptr sp Int.zero) - te1 tm1 (store_parameters cenv params) - E0 te2 tm2 Out_normal - /\ mem_inject f m2 tm2 - /\ match_callstack f (mkframe cenv e te2 sp lo hi :: cs) m2.(nextblock) tm2.(nextblock) m2. -Proof. - induction 1. - (* base case *) - intros; simpl. exists te1; exists tm1. split. constructor. tauto. - (* inductive case *) - intros until tm1. intros VVM NOREPET MINJ MATCH. simpl. - inversion VVM. subst f0 id0 chunk0 vars v vals te. - inversion MATCH. subst f0 cenv0 e0 te sp0 lo0 hi0 cs0 bound tbound m0. - inversion H18. - inversion NOREPET. subst hd tl. - assert (NEXT: nextblock m1 = nextblock m). - exploit store_inv; eauto. simpl; tauto. - generalize (me_vars0 id). intro. inversion H2; subst. - (* cenv!!id = Var_local chunk *) - assert (b0 = b). congruence. subst b0. - assert (chunk0 = chunk). congruence. subst chunk0. - assert (v' = tv). congruence. subst v'. - exploit make_cast_correct. - apply eval_Evar with (id := id). eauto. - eexact H10. - intros [tv' [EVAL1 VINJ1]]. - set (te2 := PTree.set id tv' te1). - assert (VVM2: vars_vals_match f params vl te2). - apply vars_vals_match_extensional with te1; auto. - intros. unfold te2; apply PTree.gso. red; intro; subst id0. - elim H4. change id with (fst (id, lv)). apply List.in_map; auto. - exploit store_unmapped_inject; eauto. intro MINJ2. - exploit match_callstack_store_local; eauto. - fold te2; rewrite <- NEXT; intro MATCH2. - exploit IHbind_parameters; eauto. - intros [te3 [tm3 [EXEC3 [MINJ3 MATCH3]]]]. - exists te3; exists tm3. - (* execution *) - split. apply exec_Sseq_continue with E0 te2 tm1 E0. - econstructor. unfold te2. constructor. eassumption. - assumption. traceEq. - (* meminj & match_callstack *) - tauto. - - (* cenv!!id = Var_stack_scalar *) - assert (b0 = b). congruence. subst b0. - assert (chunk0 = chunk). congruence. subst chunk0. - exploit make_store_correct. - eapply make_stackaddr_correct. - apply eval_Evar with (id := id). - eauto. 2:eauto. 2:eauto. unfold storev; eexact H0. eauto. - intros [tm2 [EVAL3 [MINJ2 NEXT1]]]. - exploit match_callstack_mapped. - eexact MATCH. 2:eauto. inversion H7. congruence. - rewrite <- NEXT; rewrite <- NEXT1; intro MATCH2. - exploit IHbind_parameters; eauto. - intros [te3 [tm3 [EVAL4 [MINJ3 MATCH3]]]]. - exists te3; exists tm3. - (* execution *) - split. apply exec_Sseq_continue with (E0**E0) te1 tm2 E0. - auto. assumption. traceEq. - (* meminj & match_callstack *) - tauto. - - (* Impossible cases on cenv!!id *) - congruence. - congruence. - congruence. -Qed. - -Lemma vars_vals_match_holds_1: - forall f params args targs, - list_norepet (List.map (@fst ident memory_chunk) params) -> - List.length params = List.length args -> - val_list_inject f args targs -> - vars_vals_match f params args - (set_params targs (List.map (@fst ident memory_chunk) params)). -Proof. - induction params; destruct args; simpl; intros; try discriminate. - constructor. - inversion H1. subst v0 vl targs. - inversion H. subst hd tl. - destruct a as [id chunk]. econstructor. - simpl. rewrite PTree.gss. reflexivity. - auto. - apply vars_vals_match_extensional - with (set_params vl' (map (@fst ident memory_chunk) params)). - eapply IHparams; eauto. - intros. simpl. apply PTree.gso. red; intro; subst id0. - elim H5. change (fst (id, chunk)) with (fst (id, lv)). - apply List.in_map; auto. -Qed. - -Lemma vars_vals_match_holds: - forall f params args targs, - List.length params = List.length args -> - val_list_inject f args targs -> - forall vars, - list_norepet (List.map (@fst ident var_kind) vars - ++ List.map (@fst ident memory_chunk) params) -> - vars_vals_match f params args - (set_locals (List.map (@fst ident var_kind) vars) - (set_params targs (List.map (@fst ident memory_chunk) params))). -Proof. - induction vars; simpl; intros. - eapply vars_vals_match_holds_1; eauto. - inversion H1. subst hd tl. - eapply vars_vals_match_extensional; eauto. - intros. apply PTree.gso. red; intro; subst id; elim H4. - apply in_or_app. right. change (fst a) with (fst (fst a, lv)). - apply List.in_map; auto. -Qed. - -Lemma bind_parameters_length: - forall e m1 params args m2, - bind_parameters e m1 params args m2 -> - List.length params = List.length args. -Proof. - induction 1; simpl; eauto. -Qed. - -(** The final result in this section: the behaviour of function entry - in the generated Cminor code (allocate stack data block and store - parameters whose address is taken) simulates what happens at function - entry in the original Csharpminor (allocate one block per local variable - and initialize the blocks corresponding to function parameters). *) - -Lemma function_entry_ok: - forall fn m e m1 lb vargs m2 f cs tm cenv sz tm1 sp tvargs, - alloc_variables empty_env m (fn_variables fn) e m1 lb -> - bind_parameters e m1 fn.(Csharpminor.fn_params) vargs m2 -> - match_callstack f cs m.(nextblock) tm.(nextblock) m -> - build_compilenv gce fn = (cenv, sz) -> - sz <= Int.max_signed -> - Mem.alloc tm 0 sz = (tm1, sp) -> - let te := - set_locals (fn_vars_names fn) (set_params tvargs (fn_params_names fn)) in - val_list_inject f vargs tvargs -> - mem_inject f m tm -> - list_norepet (fn_params_names fn ++ fn_vars_names fn) -> - exists f2, exists te2, exists tm2, - exec_stmt tge (Vptr sp Int.zero) - te tm1 (store_parameters cenv fn.(Csharpminor.fn_params)) - E0 te2 tm2 Out_normal - /\ mem_inject f2 m2 tm2 - /\ inject_incr f f2 - /\ match_callstack f2 - (mkframe cenv e te2 sp m.(nextblock) m1.(nextblock) :: cs) - m2.(nextblock) tm2.(nextblock) m2 - /\ (forall b, m.(nextblock) <= b < m1.(nextblock) <-> In b lb). -Proof. - intros. - exploit bind_parameters_length; eauto. intro LEN1. - exploit match_callstack_alloc_variables; eauto. - intros [f1 [INCR1 [MINJ1 MATCH1]]]. - exploit vars_vals_match_holds. - eauto. apply val_list_inject_incr with f. eauto. eauto. - apply list_norepet_append_commut. - unfold fn_vars_names in H7. eexact H7. - intro VVM. - exploit store_parameters_correct. - eauto. eauto. - unfold fn_params_names in H7. eapply list_norepet_append_left; eauto. - eexact MINJ1. eauto. - intros [te2 [tm2 [EXEC [MINJ2 MATCH2]]]]. - exists f1; exists te2; exists tm2. - split; auto. split; auto. split; auto. split; auto. - intros; eapply alloc_variables_list_block; eauto. -Qed. - -(** * Semantic preservation for the translation *) - -(** These tactics simplify hypotheses of the form [f ... = Some x]. *) - -Ltac monadSimpl1 := - match goal with - | [ |- (bind _ _ ?F ?G = Some ?X) -> _ ] => - unfold bind at 1; - generalize (refl_equal F); - pattern F at -1 in |- *; - case F; - [ (let EQ := fresh "EQ" in - (intro; intro EQ; - try monadSimpl1)) - | intros; discriminate ] - | [ |- (None = Some _) -> _ ] => - intro; discriminate - | [ |- (Some _ = Some _) -> _ ] => - let h := fresh "H" in - (intro h; injection h; intro; clear h) - end. - -Ltac monadSimpl := - match goal with - | [ |- (bind _ _ ?F ?G = Some ?X) -> _ ] => monadSimpl1 - | [ |- (None = Some _) -> _ ] => monadSimpl1 - | [ |- (Some _ = Some _) -> _ ] => monadSimpl1 - | [ |- (?F _ _ _ _ _ _ _ = Some _) -> _ ] => simpl F; monadSimpl1 - | [ |- (?F _ _ _ _ _ _ = Some _) -> _ ] => simpl F; monadSimpl1 - | [ |- (?F _ _ _ _ _ = Some _) -> _ ] => simpl F; monadSimpl1 - | [ |- (?F _ _ _ _ = Some _) -> _ ] => simpl F; monadSimpl1 - | [ |- (?F _ _ _ = Some _) -> _ ] => simpl F; monadSimpl1 - | [ |- (?F _ _ = Some _) -> _ ] => simpl F; monadSimpl1 - | [ |- (?F _ = Some _) -> _ ] => simpl F; monadSimpl1 - end. - -Ltac monadInv H := - generalize H; monadSimpl. - -(** The proof of semantic preservation uses simulation diagrams of the - following form: -<< - le, e, m1, a --------------- tle, sp, te1, tm1, ta - | | - | | - v v - le, e, m2, v --------------- tle, sp, te2, tm2, tv ->> - where [ta] is the Cminor expression obtained by translating the - Csharpminor expression [a]. The left vertical arrow is an evaluation - of a Csharpminor expression. The right vertical arrow is an evaluation - of a Cminor expression. The precondition (top vertical bar) - includes a [mem_inject] relation between the memory states [m1] and [tm1], - a [val_list_inject] relation between the let environments [le] and [tle], - and a [match_callstack] relation for any callstack having - [e], [te1], [sp] as top frame. The postcondition (bottom vertical bar) - is the existence of a memory injection [f2] that extends the injection - [f1] we started with, preserves the [match_callstack] relation for - the transformed callstack at the final state, and validates a - [val_inject] relation between the result values [v] and [tv]. - - We capture these diagrams by the following predicates, parameterized - over the Csharpminor executions, which will serve as induction - hypotheses in the proof of simulation. *) - -Definition eval_expr_prop - (le: Csharpminor.letenv) (e: Csharpminor.env) (m1: mem) (a: Csharpminor.expr) (t: trace) (m2: mem) (v: val) : Prop := - forall cenv ta f1 tle te1 tm1 sp lo hi cs - (TR: transl_expr cenv a = Some ta) - (LINJ: val_list_inject f1 le tle) - (MINJ: mem_inject f1 m1 tm1) - (MATCH: match_callstack f1 - (mkframe cenv e te1 sp lo hi :: cs) - m1.(nextblock) tm1.(nextblock) m1), - exists f2, exists te2, exists tm2, exists tv, - eval_expr tge (Vptr sp Int.zero) tle te1 tm1 ta t te2 tm2 tv - /\ val_inject f2 v tv - /\ mem_inject f2 m2 tm2 - /\ inject_incr f1 f2 - /\ match_callstack f2 - (mkframe cenv e te2 sp lo hi :: cs) - m2.(nextblock) tm2.(nextblock) m2. - -Definition eval_exprlist_prop - (le: Csharpminor.letenv) (e: Csharpminor.env) (m1: mem) (al: Csharpminor.exprlist) (t: trace) (m2: mem) (vl: list val) : Prop := - forall cenv tal f1 tle te1 tm1 sp lo hi cs - (TR: transl_exprlist cenv al = Some tal) - (LINJ: val_list_inject f1 le tle) - (MINJ: mem_inject f1 m1 tm1) - (MATCH: match_callstack f1 - (mkframe cenv e te1 sp lo hi :: cs) - m1.(nextblock) tm1.(nextblock) m1), - exists f2, exists te2, exists tm2, exists tvl, - eval_exprlist tge (Vptr sp Int.zero) tle te1 tm1 tal t te2 tm2 tvl - /\ val_list_inject f2 vl tvl - /\ mem_inject f2 m2 tm2 - /\ inject_incr f1 f2 - /\ match_callstack f2 - (mkframe cenv e te2 sp lo hi :: cs) - m2.(nextblock) tm2.(nextblock) m2. - -Definition eval_funcall_prop - (m1: mem) (fn: Csharpminor.fundef) (args: list val) (t: trace) (m2: mem) (res: val) : Prop := - forall tfn f1 tm1 cs targs - (TR: transl_fundef gce fn = Some tfn) - (MINJ: mem_inject f1 m1 tm1) - (MATCH: match_callstack f1 cs m1.(nextblock) tm1.(nextblock) m1) - (ARGSINJ: val_list_inject f1 args targs), - exists f2, exists tm2, exists tres, - eval_funcall tge tm1 tfn targs t tm2 tres - /\ val_inject f2 res tres - /\ mem_inject f2 m2 tm2 - /\ inject_incr f1 f2 - /\ match_callstack f2 cs m2.(nextblock) tm2.(nextblock) m2. - -Inductive outcome_inject (f: meminj) : Csharpminor.outcome -> outcome -> Prop := - | outcome_inject_normal: - outcome_inject f Csharpminor.Out_normal Out_normal - | outcome_inject_exit: - forall n, outcome_inject f (Csharpminor.Out_exit n) (Out_exit n) - | outcome_inject_return_none: - outcome_inject f (Csharpminor.Out_return None) (Out_return None) - | outcome_inject_return_some: - forall v1 v2, - val_inject f v1 v2 -> - outcome_inject f (Csharpminor.Out_return (Some v1)) (Out_return (Some v2)). - -Definition exec_stmt_prop - (e: Csharpminor.env) (m1: mem) (s: Csharpminor.stmt) (t: trace) (m2: mem) (out: Csharpminor.outcome): Prop := - forall cenv ts f1 te1 tm1 sp lo hi cs - (TR: transl_stmt cenv s = Some ts) - (MINJ: mem_inject f1 m1 tm1) - (MATCH: match_callstack f1 - (mkframe cenv e te1 sp lo hi :: cs) - m1.(nextblock) tm1.(nextblock) m1), - exists f2, exists te2, exists tm2, exists tout, - exec_stmt tge (Vptr sp Int.zero) te1 tm1 ts t te2 tm2 tout - /\ outcome_inject f2 out tout - /\ mem_inject f2 m2 tm2 - /\ inject_incr f1 f2 - /\ match_callstack f2 - (mkframe cenv e te2 sp lo hi :: cs) - m2.(nextblock) tm2.(nextblock) m2. - -(** There are as many cases in the inductive proof as there are evaluation - rules in the Csharpminor semantics. We treat each case as a separate - lemma. *) - -Lemma transl_expr_Evar_correct: - forall (le : Csharpminor.letenv) - (e : Csharpminor.env) (m : mem) (id : positive) - (b : block) (chunk : memory_chunk) (v : val), - eval_var_ref prog e id b chunk -> - load chunk m b 0 = Some v -> - eval_expr_prop le e m (Csharpminor.Evar id) E0 m v. -Proof. - intros; red; intros. unfold transl_expr in TR. - exploit var_get_correct; eauto. - intros [tv [EVAL VINJ]]. - exists f1; exists te1; exists tm1; exists tv; intuition eauto. -Qed. - -Lemma transl_expr_Eaddrof_correct: - forall (le : Csharpminor.letenv) - (e : Csharpminor.env) (m : mem) (id : positive) - (b : block), - eval_var_addr prog e id b -> - eval_expr_prop le e m (Eaddrof id) E0 m (Vptr b Int.zero). -Proof. - intros; red; intros. simpl in TR. - exploit var_addr_correct; eauto. - intros [tv [EVAL VINJ]]. - exists f1; exists te1; exists tm1; exists tv. intuition eauto. -Qed. - -Lemma transl_expr_Eop_correct: - forall (le : Csharpminor.letenv) (e : Csharpminor.env) (m : mem) - (op : Csharpminor.operation) (al : Csharpminor.exprlist) - (t: trace) (m1 : mem) (vl : list val) (v : val), - Csharpminor.eval_exprlist prog le e m al t m1 vl -> - eval_exprlist_prop le e m al t m1 vl -> - Csharpminor.eval_operation op vl m1 = Some v -> - eval_expr_prop le e m (Csharpminor.Eop op al) t m1 v. -Proof. - intros; red; intros. monadInv TR; intro EQ0. - exploit H0; eauto. - intros [f2 [te2 [tm2 [tvl [EVAL1 [VINJ1 [MINJ1 [INCR1 MATCH1]]]]]]]]. - exploit make_op_correct; eauto. - intros [tv [EVAL2 VINJ2]]. - exists f2; exists te2; exists tm2; exists tv. intuition. -Qed. - -Lemma transl_expr_Eload_correct: - forall (le : Csharpminor.letenv) (e : Csharpminor.env) (m : mem) - (chunk : memory_chunk) (a : Csharpminor.expr) (t: trace) (m1 : mem) - (v1 v : val), - Csharpminor.eval_expr prog le e m a t m1 v1 -> - eval_expr_prop le e m a t m1 v1 -> - loadv chunk m1 v1 = Some v -> - eval_expr_prop le e m (Csharpminor.Eload chunk a) t m1 v. -Proof. - intros; red; intros. - monadInv TR. - exploit H0; eauto. - intros [f2 [te2 [tm2 [tv1 [EVAL [VINJ1 [MINJ2 [INCR MATCH2]]]]]]]]. - exploit loadv_inject; eauto. - intros [tv [TLOAD VINJ]]. - exists f2; exists te2; exists tm2; exists tv. - intuition. - subst ta. eapply make_load_correct; eauto. -Qed. - -Lemma transl_expr_Ecall_correct: - forall (le : Csharpminor.letenv) (e : Csharpminor.env) (m : mem) - (sig : signature) (a : Csharpminor.expr) (bl : Csharpminor.exprlist) - (t1: trace) (m1: mem) (t2: trace) (m2: mem) (t3: trace) (m3: mem) - (vf : val) (vargs : list val) (vres : val) - (f : Csharpminor.fundef) (t: trace), - Csharpminor.eval_expr prog le e m a t1 m1 vf -> - eval_expr_prop le e m a t1 m1 vf -> - Csharpminor.eval_exprlist prog le e m1 bl t2 m2 vargs -> - eval_exprlist_prop le e m1 bl t2 m2 vargs -> - Genv.find_funct ge vf = Some f -> - Csharpminor.funsig f = sig -> - Csharpminor.eval_funcall prog m2 f vargs t3 m3 vres -> - eval_funcall_prop m2 f vargs t3 m3 vres -> - t = t1 ** t2 ** t3 -> - eval_expr_prop le e m (Csharpminor.Ecall sig a bl) t m3 vres. -Proof. - intros;red;intros. monadInv TR. subst ta. - exploit H0; eauto. - intros [f2 [te2 [tm2 [tv1 [EVAL1 [VINJ1 [MINJ1 [INCR1 MATCH1]]]]]]]]. - exploit H2. - eauto. eapply val_list_inject_incr; eauto. eauto. eauto. - intros [f3 [te3 [tm3 [tv2 [EVAL2 [VINJ2 [MINJ2 [INCR2 MATCH2]]]]]]]]. - assert (tv1 = vf). - elim (Genv.find_funct_inv H3). intros bf VF. rewrite VF in H3. - rewrite Genv.find_funct_find_funct_ptr in H3. - generalize (Genv.find_funct_ptr_inv H3). intro. - assert (match_globalenvs f2). eapply match_callstack_match_globalenvs; eauto. - generalize (mg_functions _ H9 _ H8). intro. - rewrite VF in VINJ1. inversion VINJ1. subst vf. - decEq. congruence. - subst ofs2. replace x with 0. reflexivity. congruence. - subst tv1. elim (functions_translated _ _ H3). intros tf [FIND TRF]. - exploit H6; eauto. - intros [f4 [tm4 [tres [EVAL3 [VINJ3 [MINJ3 [INCR3 MATCH3]]]]]]]. - exists f4; exists te3; exists tm4; exists tres. intuition. - eapply eval_Ecall; eauto. - rewrite <- H4. apply sig_preserved; auto. - apply inject_incr_trans with f2; auto. - apply inject_incr_trans with f3; auto. -Qed. - -Lemma transl_expr_Econdition_true_correct: - forall (le : Csharpminor.letenv) (e : Csharpminor.env) (m : mem) - (a b c : Csharpminor.expr) (t1: trace) (m1 : mem) (v1 : val) - (t2: trace) (m2 : mem) (v2 : val) (t: trace), - Csharpminor.eval_expr prog le e m a t1 m1 v1 -> - eval_expr_prop le e m a t1 m1 v1 -> - Val.is_true v1 -> - Csharpminor.eval_expr prog le e m1 b t2 m2 v2 -> - eval_expr_prop le e m1 b t2 m2 v2 -> - t = t1 ** t2 -> - eval_expr_prop le e m (Csharpminor.Econdition a b c) t m2 v2. -Proof. - intros; red; intros. monadInv TR. - exploit H0; eauto. - intros [f2 [te2 [tm2 [tv1 [EVAL1 [VINJ1 [MINJ1 [INCR1 MATCH1]]]]]]]]. - exploit H3. - eauto. eapply val_list_inject_incr; eauto. eauto. eauto. - intros [f3 [te3 [tm3 [tv2 [EVAL2 [VINJ2 [MINJ2 [INCR2 MATCH2]]]]]]]]. - exists f3; exists te3; exists tm3; exists tv2. - intuition. - rewrite <- H6. subst t; eapply eval_conditionalexpr_true; eauto. - inversion VINJ1; subst v1 tv1; simpl in H1; simpl; contradiction || auto. - eapply inject_incr_trans; eauto. -Qed. - -Lemma transl_expr_Econdition_false_correct: - forall (le : Csharpminor.letenv) (e : Csharpminor.env) (m : mem) - (a b c : Csharpminor.expr) (t1: trace) (m1 : mem) (v1 : val) - (t2: trace) (m2 : mem) (v2 : val) (t: trace), - Csharpminor.eval_expr prog le e m a t1 m1 v1 -> - eval_expr_prop le e m a t1 m1 v1 -> - Val.is_false v1 -> - Csharpminor.eval_expr prog le e m1 c t2 m2 v2 -> - eval_expr_prop le e m1 c t2 m2 v2 -> - t = t1 ** t2 -> - eval_expr_prop le e m (Csharpminor.Econdition a b c) t m2 v2. -Proof. - intros; red; intros. monadInv TR. - exploit H0; eauto. - intros [f2 [te2 [tm2 [tv1 [EVAL1 [VINJ1 [MINJ1 [INCR1 MATCH1]]]]]]]]. - exploit H3. - eauto. eapply val_list_inject_incr; eauto. eauto. eauto. - intros [f3 [te3 [tm3 [tv2 [EVAL2 [VINJ2 [MINJ2 [INCR2 MATCH2]]]]]]]]. - exists f3; exists te3; exists tm3; exists tv2. - intuition. - rewrite <- H6. subst t; eapply eval_conditionalexpr_false; eauto. - inversion VINJ1; subst v1 tv1; simpl in H1; simpl; contradiction || auto. - eapply inject_incr_trans; eauto. -Qed. - -Lemma transl_expr_Elet_correct: - forall (le : Csharpminor.letenv) (e : Csharpminor.env) (m : mem) - (a b : Csharpminor.expr) (t1: trace) (m1 : mem) (v1 : val) - (t2: trace) (m2 : mem) (v2 : val) (t: trace), - Csharpminor.eval_expr prog le e m a t1 m1 v1 -> - eval_expr_prop le e m a t1 m1 v1 -> - Csharpminor.eval_expr prog (v1 :: le) e m1 b t2 m2 v2 -> - eval_expr_prop (v1 :: le) e m1 b t2 m2 v2 -> - t = t1 ** t2 -> - eval_expr_prop le e m (Csharpminor.Elet a b) t m2 v2. -Proof. - intros; red; intros. monadInv TR. - exploit H0; eauto. - intros [f2 [te2 [tm2 [tv1 [EVAL1 [VINJ1 [MINJ1 [INCR1 MATCH1]]]]]]]]. - exploit H2. - eauto. - constructor. eauto. eapply val_list_inject_incr; eauto. - eauto. eauto. - intros [f3 [te3 [tm3 [tv2 [EVAL2 [VINJ2 [MINJ2 [INCR2 MATCH2]]]]]]]]. - exists f3; exists te3; exists tm3; exists tv2. - intuition. - subst ta; eapply eval_Elet; eauto. - eapply inject_incr_trans; eauto. -Qed. - -Remark val_list_inject_nth: - forall f l1 l2, val_list_inject f l1 l2 -> - forall n v1, nth_error l1 n = Some v1 -> - exists v2, nth_error l2 n = Some v2 /\ val_inject f v1 v2. -Proof. - induction 1; destruct n; simpl; intros. - discriminate. discriminate. - injection H1; intros; subst v. exists v'; split; auto. - eauto. -Qed. - -Lemma transl_expr_Eletvar_correct: - forall (le : list val) (e : Csharpminor.env) (m : mem) (n : nat) - (v : val), - nth_error le n = Some v -> - eval_expr_prop le e m (Csharpminor.Eletvar n) E0 m v. -Proof. - intros; red; intros. monadInv TR. - exploit val_list_inject_nth; eauto. intros [tv [A B]]. - exists f1; exists te1; exists tm1; exists tv. - intuition. - subst ta. eapply eval_Eletvar; auto. -Qed. - -Lemma transl_expr_Ealloc_correct: - forall (le: list val) (e: Csharpminor.env) (m1: mem) (a: Csharpminor.expr) - (t: trace) (m2: mem) (n: int) (m3: mem) (b: block), - Csharpminor.eval_expr prog le e m1 a t m2 (Vint n) -> - eval_expr_prop le e m1 a t m2 (Vint n) -> - Mem.alloc m2 0 (Int.signed n) = (m3, b) -> - eval_expr_prop le e m1 (Csharpminor.Ealloc a) t m3 (Vptr b Int.zero). -Proof. - intros; red; intros. monadInv TR. - exploit H0; eauto. - intros [f2 [te2 [tm2 [tv1 [EVAL1 [VINJ1 [MINJ2 [INCR2 MATCH2]]]]]]]]. - inversion VINJ1. subst tv1 i. - caseEq (alloc tm2 0 (Int.signed n)). intros tm3 tb TALLOC. - assert (LB: Int.min_signed <= 0). compute. congruence. - assert (HB: Int.signed n <= Int.max_signed). - generalize (Int.signed_range n); omega. - exploit alloc_parallel_inject; eauto. - intros [MINJ3 INCR3]. - exists (extend_inject b (Some (tb, 0)) f2); - exists te2; exists tm3; exists (Vptr tb Int.zero). - split. subst ta; econstructor; eauto. - split. econstructor. unfold extend_inject, eq_block. rewrite zeq_true. reflexivity. - reflexivity. - split. assumption. - split. eapply inject_incr_trans; eauto. - eapply match_callstack_alloc; eauto. -Qed. - -Lemma transl_exprlist_Enil_correct: - forall (le : Csharpminor.letenv) (e : Csharpminor.env) (m : mem), - eval_exprlist_prop le e m Csharpminor.Enil E0 m nil. -Proof. - intros; red; intros. monadInv TR. - exists f1; exists te1; exists tm1; exists (@nil val). - intuition. subst tal; constructor. -Qed. - -Lemma transl_exprlist_Econs_correct: - forall (le : Csharpminor.letenv) (e : Csharpminor.env) (m : mem) - (a : Csharpminor.expr) (bl : Csharpminor.exprlist) - (t1: trace) (m1 : mem) (v : val) - (t2: trace) (m2 : mem) (vl : list val) (t: trace), - Csharpminor.eval_expr prog le e m a t1 m1 v -> - eval_expr_prop le e m a t1 m1 v -> - Csharpminor.eval_exprlist prog le e m1 bl t2 m2 vl -> - eval_exprlist_prop le e m1 bl t2 m2 vl -> - t = t1 ** t2 -> - eval_exprlist_prop le e m (Csharpminor.Econs a bl) t m2 (v :: vl). -Proof. - intros; red; intros. monadInv TR. - exploit H0; eauto. - intros [f2 [te2 [tm2 [tv1 [EVAL1 [VINJ1 [MINJ2 [INCR2 MATCH2]]]]]]]]. - exploit H2. - eauto. eapply val_list_inject_incr; eauto. eauto. eauto. - intros [f3 [te3 [tm3 [tv2 [EVAL2 [VINJ2 [MINJ3 [INCR3 MATCH3]]]]]]]]. - exists f3; exists te3; exists tm3; exists (tv1 :: tv2). - intuition. subst tal; econstructor; eauto. - constructor. eapply val_inject_incr; eauto. auto. - eapply inject_incr_trans; eauto. -Qed. - -Lemma transl_funcall_internal_correct: - forall (m : mem) (f : Csharpminor.function) (vargs : list val) - (e : Csharpminor.env) (m1 : mem) (lb : list block) (m2: mem) - (t: trace) (m3 : mem) (out : Csharpminor.outcome) (vres : val), - list_norepet (fn_params_names f ++ fn_vars_names f) -> - alloc_variables empty_env m (fn_variables f) e m1 lb -> - bind_parameters e m1 (Csharpminor.fn_params f) vargs m2 -> - Csharpminor.exec_stmt prog e m2 (Csharpminor.fn_body f) t m3 out -> - exec_stmt_prop e m2 (Csharpminor.fn_body f) t m3 out -> - Csharpminor.outcome_result_value out (sig_res (Csharpminor.fn_sig f)) vres -> - eval_funcall_prop m (Internal f) vargs t (free_list m3 lb) vres. -Proof. - intros; red. intros tfn f1 tm; intros. - generalize TR; clear TR. - unfold transl_fundef, transf_partial_fundef. - caseEq (transl_function gce f); try congruence. - intros tf TR EQ. inversion EQ; clear EQ; subst tfn. - unfold transl_function in TR. - caseEq (build_compilenv gce f); intros cenv stacksize CENV. - rewrite CENV in TR. - destruct (zle stacksize Int.max_signed); try discriminate. - monadInv TR. clear TR. - caseEq (alloc tm 0 stacksize). intros tm1 sp ALLOC. - exploit function_entry_ok; eauto. - intros [f2 [te2 [tm2 [STOREPARAM [MINJ2 [INCR12 [MATCH2 BLOCKS]]]]]]]. - red in H3; exploit H3; eauto. - intros [f3 [te3 [tm3 [tout [EXECBODY [OUTINJ [MINJ3 [INCR23 MATCH3]]]]]]]]. - assert (exists tvres, - outcome_result_value tout f.(Csharpminor.fn_sig).(sig_res) tvres /\ - val_inject f3 vres tvres). - generalize H4. unfold Csharpminor.outcome_result_value, outcome_result_value. - inversion OUTINJ. - destruct (sig_res (Csharpminor.fn_sig f)); intro. contradiction. - exists Vundef; split. auto. subst vres; constructor. - tauto. - destruct (sig_res (Csharpminor.fn_sig f)); intro. contradiction. - exists Vundef; split. auto. subst vres; constructor. - destruct (sig_res (Csharpminor.fn_sig f)); intro. - exists v2; split. auto. subst vres; auto. - contradiction. - destruct H5 as [tvres [TOUT VINJRES]]. - exists f3; exists (Mem.free tm3 sp); exists tvres. - (* execution *) - split. rewrite <- H6; econstructor; simpl; eauto. - apply exec_Sseq_continue with E0 te2 tm2 t. - exact STOREPARAM. - eexact EXECBODY. - traceEq. - (* val_inject *) - split. assumption. - (* mem_inject *) - split. apply free_inject; auto. - intros. elim (BLOCKS b1); intros B1 B2. apply B1. inversion MATCH3. inversion H20. - eapply me_inv0. eauto. - (* inject_incr *) - split. eapply inject_incr_trans; eauto. - (* match_callstack *) - assert (forall bl mm, nextblock (free_list mm bl) = nextblock mm). - induction bl; intros. reflexivity. simpl. auto. - unfold free; simpl nextblock. rewrite H5. - eapply match_callstack_freelist; eauto. - intros. elim (BLOCKS b); intros B1 B2. generalize (B2 H7). omega. -Qed. - -Lemma transl_funcall_external_correct: - forall (m : mem) (ef : external_function) (vargs : list val) - (t : trace) (vres : val), - event_match ef vargs t vres -> - eval_funcall_prop m (External ef) vargs t m vres. -Proof. - intros; red; intros. - simpl in TR. inversion TR; clear TR; subst tfn. - exploit event_match_inject; eauto. intros [A B]. - exists f1; exists tm1; exists vres; intuition. - constructor; auto. -Qed. - -Lemma transl_stmt_Sskip_correct: - forall (e : Csharpminor.env) (m : mem), - exec_stmt_prop e m Csharpminor.Sskip E0 m Csharpminor.Out_normal. -Proof. - intros; red; intros. monadInv TR. - exists f1; exists te1; exists tm1; exists Out_normal. - intuition. subst ts; constructor. constructor. -Qed. - -Lemma transl_stmt_Sexpr_correct: - forall (e : Csharpminor.env) (m : mem) (a : Csharpminor.expr) - (t: trace) (m1 : mem) (v : val), - Csharpminor.eval_expr prog nil e m a t m1 v -> - eval_expr_prop nil e m a t m1 v -> - exec_stmt_prop e m (Csharpminor.Sexpr a) t m1 Csharpminor.Out_normal. -Proof. - intros; red; intros. monadInv TR. - exploit H0; eauto. - intros [f2 [te2 [tm2 [tv1 [EVAL1 [VINJ1 [MINJ2 [INCR2 MATCH2]]]]]]]]. - exists f2; exists te2; exists tm2; exists Out_normal. - intuition. subst ts. econstructor; eauto. - constructor. -Qed. - -Lemma transl_stmt_Sassign_correct: - forall (e : Csharpminor.env) (m : mem) - (id : ident) (a : Csharpminor.expr) (t: trace) (m1 : mem) (b : block) - (chunk : memory_chunk) (v : val) (m2 : mem), - Csharpminor.eval_expr prog nil e m a t m1 v -> - eval_expr_prop nil e m a t m1 v -> - eval_var_ref prog e id b chunk -> - store chunk m1 b 0 v = Some m2 -> - exec_stmt_prop e m (Csharpminor.Sassign id a) t m2 Csharpminor.Out_normal. -Proof. - intros; red; intros. monadInv TR; intro EQ0. - exploit H0; eauto. - intros [f2 [te2 [tm2 [tv1 [EVAL1 [VINJ1 [MINJ1 [INCR12 MATCH1]]]]]]]]. - exploit var_set_correct; eauto. - intros [te3 [tm3 [EVAL2 [MINJ2 MATCH2]]]]. - exists f2; exists te3; exists tm3; exists Out_normal. - intuition. constructor. -Qed. - -Lemma transl_stmt_Sstore_correct: - forall (e : Csharpminor.env) (m : mem) - (chunk : memory_chunk) (a b : Csharpminor.expr) (t1: trace) (m1 : mem) - (v1 : val) (t2: trace) (m2 : mem) (v2 : val) - (t3: trace) (m3 : mem), - Csharpminor.eval_expr prog nil e m a t1 m1 v1 -> - eval_expr_prop nil e m a t1 m1 v1 -> - Csharpminor.eval_expr prog nil e m1 b t2 m2 v2 -> - eval_expr_prop nil e m1 b t2 m2 v2 -> - storev chunk m2 v1 v2 = Some m3 -> - t3 = t1 ** t2 -> - exec_stmt_prop e m (Csharpminor.Sstore chunk a b) t3 m3 Csharpminor.Out_normal. -Proof. - intros; red; intros. monadInv TR. - exploit H0; eauto. - intros [f2 [te2 [tm2 [tv1 [EVAL1 [VINJ1 [MINJ2 [INCR2 MATCH2]]]]]]]]. - exploit H2. - eauto. - eapply val_list_inject_incr; eauto. - eauto. eauto. - intros [f3 [te3 [tm3 [tv2 [EVAL2 [VINJ2 [MINJ3 [INCR3 MATCH3]]]]]]]]. - exploit make_store_correct. - eexact EVAL1. eexact EVAL2. eauto. eauto. - eapply val_inject_incr; eauto. eauto. - intros [tm4 [EVAL [MINJ4 NEXTBLOCK]]]. - exists f3; exists te3; exists tm4; exists Out_normal. - rewrite <- H6. subst t3. intuition. - constructor. - eapply inject_incr_trans; eauto. - assert (val_inject f3 v1 tv1). eapply val_inject_incr; eauto. - unfold storev in H3; destruct v1; try discriminate. - inversion H4. - rewrite NEXTBLOCK. replace (nextblock m3) with (nextblock m2). - eapply match_callstack_mapped; eauto. congruence. - exploit store_inv; eauto. simpl; symmetry; tauto. -Qed. - -Lemma transl_stmt_Sseq_continue_correct: - forall (e : Csharpminor.env) (m : mem) (s1 s2 : Csharpminor.stmt) - (t1 t2: trace) (m1 m2 : mem) (t: trace) (out : Csharpminor.outcome), - Csharpminor.exec_stmt prog e m s1 t1 m1 Csharpminor.Out_normal -> - exec_stmt_prop e m s1 t1 m1 Csharpminor.Out_normal -> - Csharpminor.exec_stmt prog e m1 s2 t2 m2 out -> - exec_stmt_prop e m1 s2 t2 m2 out -> - t = t1 ** t2 -> - exec_stmt_prop e m (Csharpminor.Sseq s1 s2) t m2 out. -Proof. - intros; red; intros; monadInv TR. - exploit H0; eauto. - intros [f2 [te2 [tm2 [tout1 [EXEC1 [OINJ1 [MINJ2 [INCR2 MATCH2]]]]]]]]. - exploit H2; eauto. - intros [f3 [te3 [tm3 [tout2 [EXEC2 [OINJ2 [MINJ3 [INCR3 MATCH3]]]]]]]]. - exists f3; exists te3; exists tm3; exists tout2. - intuition. subst ts; eapply exec_Sseq_continue; eauto. - inversion OINJ1. subst tout1. auto. - eapply inject_incr_trans; eauto. -Qed. - -Lemma transl_stmt_Sseq_stop_correct: - forall (e : Csharpminor.env) (m : mem) (s1 s2 : Csharpminor.stmt) - (t1: trace) (m1 : mem) (out : Csharpminor.outcome), - Csharpminor.exec_stmt prog e m s1 t1 m1 out -> - exec_stmt_prop e m s1 t1 m1 out -> - out <> Csharpminor.Out_normal -> - exec_stmt_prop e m (Csharpminor.Sseq s1 s2) t1 m1 out. -Proof. - intros; red; intros; monadInv TR. - exploit H0; eauto. - intros [f2 [te2 [tm2 [tout1 [EXEC1 [OINJ1 [MINJ2 [INCR2 MATCH2]]]]]]]]. - exists f2; exists te2; exists tm2; exists tout1. - intuition. subst ts; eapply exec_Sseq_stop; eauto. - inversion OINJ1; subst out tout1; congruence. -Qed. - -Lemma transl_stmt_Sifthenelse_true_correct: - forall (e : Csharpminor.env) (m : mem) (a : Csharpminor.expr) - (sl1 sl2 : Csharpminor.stmt) - (t1: trace) (m1 : mem) (v1 : val) (t2: trace) (m2 : mem) - (out : Csharpminor.outcome) (t: trace), - Csharpminor.eval_expr prog nil e m a t1 m1 v1 -> - eval_expr_prop nil e m a t1 m1 v1 -> - Val.is_true v1 -> - Csharpminor.exec_stmt prog e m1 sl1 t2 m2 out -> - exec_stmt_prop e m1 sl1 t2 m2 out -> - t = t1 ** t2 -> - exec_stmt_prop e m (Csharpminor.Sifthenelse a sl1 sl2) t m2 out. -Proof. - intros; red; intros. monadInv TR. - exploit H0; eauto. - intros [f2 [te2 [tm2 [tv1 [EVAL1 [VINJ1 [MINJ2 [INCR2 MATCH2]]]]]]]]. - exploit H3; eauto. - intros [f3 [te3 [tm3 [tout [EVAL2 [OINJ [MINJ3 [INCR3 MATCH3]]]]]]]]. - exists f3; exists te3; exists tm3; exists tout. - intuition. - subst ts t. eapply exec_ifthenelse_true; eauto. - inversion VINJ1; subst v1 tv1; simpl in H1; simpl; contradiction || auto. - eapply inject_incr_trans; eauto. -Qed. - -Lemma transl_stmt_Sifthenelse_false_correct: - forall (e : Csharpminor.env) (m : mem) (a : Csharpminor.expr) - (sl1 sl2 : Csharpminor.stmt) - (t1: trace) (m1 : mem) (v1 : val) (t2: trace) (m2 : mem) - (out : Csharpminor.outcome) (t: trace), - Csharpminor.eval_expr prog nil e m a t1 m1 v1 -> - eval_expr_prop nil e m a t1 m1 v1 -> - Val.is_false v1 -> - Csharpminor.exec_stmt prog e m1 sl2 t2 m2 out -> - exec_stmt_prop e m1 sl2 t2 m2 out -> - t = t1 ** t2 -> - exec_stmt_prop e m (Csharpminor.Sifthenelse a sl1 sl2) t m2 out. -Proof. - intros; red; intros. monadInv TR. - exploit H0; eauto. - intros [f2 [te2 [tm2 [tv1 [EVAL1 [VINJ1 [MINJ2 [INCR2 MATCH2]]]]]]]]. - exploit H3; eauto. - intros [f3 [te3 [tm3 [tout [EVAL2 [OINJ [MINJ3 [INCR3 MATCH3]]]]]]]]. - exists f3; exists te3; exists tm3; exists tout. - intuition. - subst ts t. eapply exec_ifthenelse_false; eauto. - inversion VINJ1; subst v1 tv1; simpl in H1; simpl; contradiction || auto. - eapply inject_incr_trans; eauto. -Qed. - -Lemma transl_stmt_Sloop_loop_correct: - forall (e : Csharpminor.env) (m : mem) (sl : Csharpminor.stmt) - (t1: trace) (m1: mem) (t2: trace) (m2 : mem) - (out : Csharpminor.outcome) (t: trace), - Csharpminor.exec_stmt prog e m sl t1 m1 Csharpminor.Out_normal -> - exec_stmt_prop e m sl t1 m1 Csharpminor.Out_normal -> - Csharpminor.exec_stmt prog e m1 (Csharpminor.Sloop sl) t2 m2 out -> - exec_stmt_prop e m1 (Csharpminor.Sloop sl) t2 m2 out -> - t = t1 ** t2 -> - exec_stmt_prop e m (Csharpminor.Sloop sl) t m2 out. -Proof. - intros; red; intros. monadInv TR. - exploit H0; eauto. - intros [f2 [te2 [tm2 [tout1 [EVAL1 [OINJ1 [MINJ2 [INCR2 MATCH2]]]]]]]]. - exploit H2; eauto. - intros [f3 [te3 [tm3 [tout2 [EVAL2 [OINJ2 [MINJ3 [INCR3 MATCH3]]]]]]]]. - exists f3; exists te3; exists tm3; exists tout2. - intuition. - subst ts. eapply exec_Sloop_loop; eauto. - inversion OINJ1; subst tout1; eauto. - eapply inject_incr_trans; eauto. -Qed. - -Lemma transl_stmt_Sloop_exit_correct: - forall (e : Csharpminor.env) (m : mem) (sl : Csharpminor.stmt) - (t1: trace) (m1 : mem) (out : Csharpminor.outcome), - Csharpminor.exec_stmt prog e m sl t1 m1 out -> - exec_stmt_prop e m sl t1 m1 out -> - out <> Csharpminor.Out_normal -> - exec_stmt_prop e m (Csharpminor.Sloop sl) t1 m1 out. -Proof. - intros; red; intros. monadInv TR. - exploit H0; eauto. - intros [f2 [te2 [tm2 [tout1 [EVAL1 [OINJ1 [MINJ2 [INCR2 MATCH2]]]]]]]]. - exists f2; exists te2; exists tm2; exists tout1. - intuition. subst ts; eapply exec_Sloop_stop; eauto. - inversion OINJ1; subst out tout1; congruence. -Qed. - -Lemma transl_stmt_Sblock_correct: - forall (e : Csharpminor.env) (m : mem) (sl : Csharpminor.stmt) - (t1: trace) (m1 : mem) (out : Csharpminor.outcome), - Csharpminor.exec_stmt prog e m sl t1 m1 out -> - exec_stmt_prop e m sl t1 m1 out -> - exec_stmt_prop e m (Csharpminor.Sblock sl) t1 m1 - (Csharpminor.outcome_block out). -Proof. - intros; red; intros. monadInv TR. - exploit H0; eauto. - intros [f2 [te2 [tm2 [tout1 [EVAL1 [OINJ1 [MINJ2 [INCR2 MATCH2]]]]]]]]. - exists f2; exists te2; exists tm2; exists (outcome_block tout1). - intuition. subst ts; eapply exec_Sblock; eauto. - inversion OINJ1; subst out tout1; simpl. - constructor. - destruct n; constructor. - constructor. - constructor; auto. -Qed. - -Lemma transl_stmt_Sexit_correct: - forall (e : Csharpminor.env) (m : mem) (n : nat), - exec_stmt_prop e m (Csharpminor.Sexit n) E0 m (Csharpminor.Out_exit n). -Proof. - intros; red; intros. monadInv TR. - exists f1; exists te1; exists tm1; exists (Out_exit n). - intuition. subst ts; constructor. constructor. -Qed. - -Lemma transl_stmt_Sswitch_correct: - forall (e : Csharpminor.env) (m : mem) - (a : Csharpminor.expr) (cases : list (int * nat)) (default : nat) - (t1 : trace) (m1 : mem) (n : int), - Csharpminor.eval_expr prog nil e m a t1 m1 (Vint n) -> - eval_expr_prop nil e m a t1 m1 (Vint n) -> - exec_stmt_prop e m (Csharpminor.Sswitch a cases default) t1 m1 - (Csharpminor.Out_exit (Csharpminor.switch_target n default cases)). -Proof. - intros; red; intros. monadInv TR. - exploit H0; eauto. - intros [f2 [te2 [tm2 [tv1 [EVAL [VINJ1 [MINJ2 [INCR MATCH2]]]]]]]]. - exists f2; exists te2; exists tm2; - exists (Out_exit (switch_target n default cases)). intuition. - subst ts. constructor. inversion VINJ1. subst tv1. assumption. - constructor. -Qed. - -Lemma transl_stmt_Sreturn_none_correct: - forall (e : Csharpminor.env) (m : mem), - exec_stmt_prop e m (Csharpminor.Sreturn None) E0 m - (Csharpminor.Out_return None). -Proof. - intros; red; intros. monadInv TR. - exists f1; exists te1; exists tm1; exists (Out_return None). - intuition. subst ts; constructor. constructor. -Qed. - -Lemma transl_stmt_Sreturn_some_correct: - forall (e : Csharpminor.env) (m : mem) (a : Csharpminor.expr) - (t1: trace) (m1 : mem) (v : val), - Csharpminor.eval_expr prog nil e m a t1 m1 v -> - eval_expr_prop nil e m a t1 m1 v -> - exec_stmt_prop e m (Csharpminor.Sreturn (Some a)) t1 m1 - (Csharpminor.Out_return (Some v)). -Proof. - intros; red; intros; monadInv TR. - exploit H0; eauto. - intros [f2 [te2 [tm2 [tv1 [EVAL [VINJ1 [MINJ2 [INCR MATCH2]]]]]]]]. - exists f2; exists te2; exists tm2; exists (Out_return (Some tv1)). - intuition. subst ts; econstructor; eauto. constructor; auto. -Qed. - -(** We conclude by an induction over the structure of the Csharpminor - evaluation derivation, using the lemmas above for each case. *) - -Lemma transl_function_correct: - forall m1 f vargs t m2 vres, - Csharpminor.eval_funcall prog m1 f vargs t m2 vres -> - eval_funcall_prop m1 f vargs t m2 vres. -Proof - (eval_funcall_ind4 prog - eval_expr_prop - eval_exprlist_prop - eval_funcall_prop - exec_stmt_prop - - transl_expr_Evar_correct - transl_expr_Eaddrof_correct - transl_expr_Eop_correct - transl_expr_Eload_correct - transl_expr_Ecall_correct - transl_expr_Econdition_true_correct - transl_expr_Econdition_false_correct - transl_expr_Elet_correct - transl_expr_Eletvar_correct - transl_expr_Ealloc_correct - transl_exprlist_Enil_correct - transl_exprlist_Econs_correct - transl_funcall_internal_correct - transl_funcall_external_correct - transl_stmt_Sskip_correct - transl_stmt_Sexpr_correct - transl_stmt_Sassign_correct - transl_stmt_Sstore_correct - transl_stmt_Sseq_continue_correct - transl_stmt_Sseq_stop_correct - transl_stmt_Sifthenelse_true_correct - transl_stmt_Sifthenelse_false_correct - transl_stmt_Sloop_loop_correct - transl_stmt_Sloop_exit_correct - transl_stmt_Sblock_correct - transl_stmt_Sexit_correct - transl_stmt_Sswitch_correct - transl_stmt_Sreturn_none_correct - transl_stmt_Sreturn_some_correct). - -(** The [match_globalenvs] relation holds for the global environments - of the source program and the transformed program. *) - -Lemma match_globalenvs_init: - let m := Genv.init_mem (program_of_program prog) in - let tm := Genv.init_mem tprog in - let f := fun b => if zlt b m.(nextblock) then Some(b, 0) else None in - match_globalenvs f. -Proof. - intros. constructor. - (* globalvars *) - (* symbols *) - intros. split. - unfold f. rewrite zlt_true. auto. unfold m. - eapply Genv.find_symbol_inv; eauto. - rewrite <- H. apply symbols_preserved. - intros. unfold f. rewrite zlt_true. auto. - generalize (nextblock_pos m). omega. -Qed. - -(** The correctness of the translation of a whole Csharpminor program - follows. *) - -Theorem transl_program_correct: - forall t n, - Csharpminor.exec_program prog t (Vint n) -> - exec_program tprog t (Vint n). -Proof. - intros t n [b [fn [m [FINDS [FINDF [SIG EVAL]]]]]]. - elim (function_ptr_translated _ _ FINDF). intros tfn [TFIND TR]. - set (m0 := Genv.init_mem (program_of_program prog)) in *. - set (f := fun b => if zlt b m0.(nextblock) then Some(b, 0) else None). - assert (MINJ0: mem_inject f m0 m0). - unfold f; constructor; intros. - apply zlt_false; auto. - destruct (zlt b0 (nextblock m0)); try discriminate. - inversion H; subst b' delta. - split. auto. - constructor. compute. split; congruence. left; auto. - intros; omega. - generalize (Genv.initmem_block_init (program_of_program prog) b0). fold m0. intros [idata EQ]. - rewrite EQ. simpl. apply Mem.contents_init_data_inject. - destruct (zlt b1 (nextblock m0)); try discriminate. - destruct (zlt b2 (nextblock m0)); try discriminate. - left; congruence. - assert (MATCH0: match_callstack f nil m0.(nextblock) m0.(nextblock) m0). - constructor. unfold f; apply match_globalenvs_init. - fold ge in EVAL. - exploit transl_function_correct; eauto. - intros [f1 [tm1 [tres [TEVAL [VINJ [MINJ1 [INCR MATCH1]]]]]]]. - exists b; exists tfn; exists tm1. - split. fold tge. rewrite <- FINDS. - replace (prog_main prog) with (AST.prog_main tprog). fold ge. apply symbols_preserved. - transitivity (AST.prog_main (program_of_program prog)). - apply transform_partial_program_main with (transl_fundef gce). assumption. - reflexivity. - split. assumption. - split. rewrite <- SIG. apply sig_preserved; auto. - rewrite (Genv.init_mem_transf_partial (transl_fundef gce) _ TRANSL). - inversion VINJ; subst tres. assumption. -Qed. - -End TRANSLATION. diff --git a/backend/Csharpminor.v b/backend/Csharpminor.v deleted file mode 100644 index 246ebf53..00000000 --- a/backend/Csharpminor.v +++ /dev/null @@ -1,563 +0,0 @@ -(** Abstract syntax and semantics for the Csharpminor language. *) - -Require Import Coqlib. -Require Import Maps. -Require Import AST. -Require Import Integers. -Require Import Floats. -Require Import Values. -Require Import Mem. -Require Import Events. -Require Import Globalenvs. - -(** Abstract syntax *) - -(** Csharpminor is a low-level imperative language structured in expressions, - statements, functions and programs. Expressions include - reading global or local variables, reading store locations, - arithmetic operations, function calls, and conditional expressions - (similar to [e1 ? e2 : e3] in C). The [Elet] and [Eletvar] constructs - enable sharing the computations of subexpressions. De Bruijn notation - is used: [Eletvar n] refers to the value bound by then [n+1]-th enclosing - [Elet] construct. - - Unlike in Cminor (the next intermediate language of the back-end), - Csharpminor local variables reside in memory, and their address can - be taken using [Eaddrof] expressions. - - Another difference with Cminor is that Csharpminor is entirely - processor-neutral. In particular, Csharpminor uses a standard set of - operations: it does not reflect processor-specific operations nor - addressing modes. *) - -Inductive operation : Set := - | Ointconst: int -> operation (**r integer constant *) - | Ofloatconst: float -> operation (**r floating-point constant *) - | Ocast8unsigned: operation (**r 8-bit zero extension *) - | Ocast8signed: operation (**r 8-bit sign extension *) - | Ocast16unsigned: operation (**r 16-bit zero extension *) - | Ocast16signed: operation (**r 16-bit sign extension *) - | Onotint: operation (**r bitwise complement *) - | Oadd: operation (**r integer addition *) - | Osub: operation (**r integer subtraction *) - | Omul: operation (**r integer multiplication *) - | Odiv: operation (**r integer signed division *) - | Odivu: operation (**r integer unsigned division *) - | Omod: operation (**r integer signed modulus *) - | Omodu: operation (**r integer unsigned modulus *) - | Oand: operation (**r bitwise ``and'' *) - | Oor: operation (**r bitwise ``or'' *) - | Oxor: operation (**r bitwise ``xor'' *) - | Oshl: operation (**r left shift *) - | Oshr: operation (**r right signed shift *) - | Oshru: operation (**r right unsigned shift *) - | Onegf: operation (**r float opposite *) - | Oabsf: operation (**r float absolute value *) - | Oaddf: operation (**r float addition *) - | Osubf: operation (**r float subtraction *) - | Omulf: operation (**r float multiplication *) - | Odivf: operation (**r float division *) - | Osingleoffloat: operation (**r float truncation *) - | Ointoffloat: operation (**r integer to float *) - | Ofloatofint: operation (**r float to signed integer *) - | Ofloatofintu: operation (**r float to unsigned integer *) - | Ocmp: comparison -> operation (**r integer signed comparison *) - | Ocmpu: comparison -> operation (**r integer unsigned comparison *) - | Ocmpf: comparison -> operation. (**r float comparison *) - -Inductive expr : Set := - | Evar : ident -> expr (**r reading a scalar variable *) - | Eaddrof : ident -> expr (**r taking the address of a variable *) - | Eop : operation -> exprlist -> expr (**r arithmetic operation *) - | Eload : memory_chunk -> expr -> expr (**r memory read *) - | Ecall : signature -> expr -> exprlist -> expr (**r function call *) - | Econdition : expr -> expr -> expr -> expr (**r conditional expression *) - | Elet : expr -> expr -> expr (**r let binding *) - | Eletvar : nat -> expr (**r reference to a let-bound variable *) - | Ealloc : expr -> expr (**r memory allocation *) - -with exprlist : Set := - | Enil: exprlist - | Econs: expr -> exprlist -> exprlist. - -(** Statements include expression evaluation, variable assignment, - memory stores, an if/then/else conditional, - infinite loops, blocks and early block exits, and early function returns. - [Sexit n] terminates prematurely the execution of the [n+1] enclosing - [Sblock] statements. *) - -Inductive stmt : Set := - | Sskip: stmt - | Sexpr: expr -> stmt - | Sassign : ident -> expr -> stmt - | Sstore : memory_chunk -> expr -> expr -> stmt - | Sseq: stmt -> stmt -> stmt - | Sifthenelse: expr -> stmt -> stmt -> stmt - | Sloop: stmt -> stmt - | Sblock: stmt -> stmt - | Sexit: nat -> stmt - | Sswitch: expr -> list (int * nat) -> nat -> stmt - | Sreturn: option expr -> stmt. - -(** The variables can be either scalar variables - (whose type, size and signedness are given by a [memory_chunk] - or array variables (of the indicated sizes). The only operation - permitted on an array variable is taking its address. *) - -Inductive var_kind : Set := - | Vscalar: memory_chunk -> var_kind - | Varray: Z -> var_kind. - -(** Functions are composed of a signature, a list of parameter names - with associated memory chunks (parameters must be scalar), a list of - local variables with associated [var_kind] description, and a - statement representing the function body. *) - -Record function : Set := mkfunction { - fn_sig: signature; - fn_params: list (ident * memory_chunk); - fn_vars: list (ident * var_kind); - fn_body: stmt -}. - -Definition fundef := AST.fundef function. - -Record program : Set := mkprogram { - prog_funct: list (ident * fundef); - prog_main: ident; - prog_vars: list (ident * var_kind * list init_data) -}. - -Definition funsig (fd: fundef) := - match fd with - | Internal f => f.(fn_sig) - | External ef => ef.(ef_sig) - end. - -(** * Operational semantics *) - -(** The operational semantics for Csharpminor is given in big-step operational - style. Expressions evaluate to values, and statements evaluate to - ``outcomes'' indicating how execution should proceed afterwards. *) - -Inductive outcome: Set := - | Out_normal: outcome (**r continue in sequence *) - | Out_exit: nat -> outcome (**r terminate [n+1] enclosing blocks *) - | Out_return: option val -> outcome. (**r return immediately to caller *) - -Definition outcome_result_value - (out: outcome) (ot: option typ) (v: val) : Prop := - match out, ot with - | Out_normal, None => v = Vundef - | Out_return None, None => v = Vundef - | Out_return (Some v'), Some ty => v = v' - | _, _ => False - end. - -Definition outcome_block (out: outcome) : outcome := - match out with - | Out_normal => Out_normal - | Out_exit O => Out_normal - | Out_exit (S n) => Out_exit n - | Out_return optv => Out_return optv - end. - -Fixpoint switch_target (n: int) (dfl: nat) (cases: list (int * nat)) - {struct cases} : nat := - match cases with - | nil => dfl - | (n1, lbl1) :: rem => if Int.eq n n1 then lbl1 else switch_target n dfl rem - end. - -(** Four kinds of evaluation environments are involved: -- [genv]: global environments, define symbols and functions; -- [gvarenv]: map global variables to var info; -- [env]: local environments, map local variables to memory blocks + var info; -- [lenv]: let environments, map de Bruijn indices to values. -*) -Definition genv := Genv.t fundef. -Definition gvarenv := PTree.t var_kind. -Definition env := PTree.t (block * var_kind). -Definition empty_env : env := PTree.empty (block * var_kind). -Definition letenv := list val. - -Definition sizeof (lv: var_kind) : Z := - match lv with - | Vscalar chunk => size_chunk chunk - | Varray sz => Zmax 0 sz - end. - -Definition program_of_program (p: program): AST.program fundef := - AST.mkprogram - p.(prog_funct) - p.(prog_main) - (List.map (fun x => match x with (id, k, init) => (id, init) end) p.(prog_vars)). - -Definition fn_variables (f: function) := - List.map - (fun id_chunk => (fst id_chunk, Vscalar (snd id_chunk))) f.(fn_params) - ++ f.(fn_vars). - -Definition fn_params_names (f: function) := - List.map (@fst ident memory_chunk) f.(fn_params). - -Definition fn_vars_names (f: function) := - List.map (@fst ident var_kind) f.(fn_vars). - -Definition global_var_env (p: program): gvarenv := - List.fold_left - (fun gve x => match x with (id, k, init) => PTree.set id k gve end) - p.(prog_vars) (PTree.empty var_kind). - -(** Evaluation of operator applications. *) - -Definition eval_compare_null (c: comparison) (n: int) : option val := - if Int.eq n Int.zero - then match c with Ceq => Some Vfalse | Cne => Some Vtrue | _ => None end - else None. - -Definition eval_operation (op: operation) (vl: list val) (m: mem): option val := - match op, vl with - | Ointconst n, nil => Some (Vint n) - | Ofloatconst n, nil => Some (Vfloat n) - | Ocast8unsigned, Vint n1 :: nil => Some (Vint (Int.cast8unsigned n1)) - | Ocast8signed, Vint n1 :: nil => Some (Vint (Int.cast8signed n1)) - | Ocast16unsigned, Vint n1 :: nil => Some (Vint (Int.cast16unsigned n1)) - | Ocast16signed, Vint n1 :: nil => Some (Vint (Int.cast16signed n1)) - | Onotint, Vint n1 :: nil => Some (Vint (Int.not n1)) - | Oadd, Vint n1 :: Vint n2 :: nil => Some (Vint (Int.add n1 n2)) - | Oadd, Vint n1 :: Vptr b2 n2 :: nil => Some (Vptr b2 (Int.add n2 n1)) - | Oadd, Vptr b1 n1 :: Vint n2 :: nil => Some (Vptr b1 (Int.add n1 n2)) - | Osub, Vint n1 :: Vint n2 :: nil => Some (Vint (Int.sub n1 n2)) - | Osub, Vptr b1 n1 :: Vint n2 :: nil => Some (Vptr b1 (Int.sub n1 n2)) - | Osub, Vptr b1 n1 :: Vptr b2 n2 :: nil => - if eq_block b1 b2 then Some (Vint (Int.sub n1 n2)) else None - | Omul, Vint n1 :: Vint n2 :: nil => Some (Vint (Int.mul n1 n2)) - | Odiv, Vint n1 :: Vint n2 :: nil => - if Int.eq n2 Int.zero then None else Some (Vint (Int.divs n1 n2)) - | Odivu, Vint n1 :: Vint n2 :: nil => - if Int.eq n2 Int.zero then None else Some (Vint (Int.divu n1 n2)) - | Omod, Vint n1 :: Vint n2 :: nil => - if Int.eq n2 Int.zero then None else Some (Vint (Int.mods n1 n2)) - | Omodu, Vint n1 :: Vint n2 :: nil => - if Int.eq n2 Int.zero then None else Some (Vint (Int.modu n1 n2)) - | Oand, Vint n1 :: Vint n2 :: nil => Some (Vint (Int.and n1 n2)) - | Oor, Vint n1 :: Vint n2 :: nil => Some (Vint (Int.or n1 n2)) - | Oxor, Vint n1 :: Vint n2 :: nil => Some (Vint (Int.xor n1 n2)) - | Oshl, Vint n1 :: Vint n2 :: nil => - if Int.ltu n2 (Int.repr 32) then Some (Vint (Int.shl n1 n2)) else None - | Oshr, Vint n1 :: Vint n2 :: nil => - if Int.ltu n2 (Int.repr 32) then Some (Vint (Int.shr n1 n2)) else None - | Oshru, Vint n1 :: Vint n2 :: nil => - if Int.ltu n2 (Int.repr 32) then Some (Vint (Int.shru n1 n2)) else None - | Onegf, Vfloat f1 :: nil => Some (Vfloat (Float.neg f1)) - | Oabsf, Vfloat f1 :: nil => Some (Vfloat (Float.abs f1)) - | Oaddf, Vfloat f1 :: Vfloat f2 :: nil => Some (Vfloat (Float.add f1 f2)) - | Osubf, Vfloat f1 :: Vfloat f2 :: nil => Some (Vfloat (Float.sub f1 f2)) - | Omulf, Vfloat f1 :: Vfloat f2 :: nil => Some (Vfloat (Float.mul f1 f2)) - | Odivf, Vfloat f1 :: Vfloat f2 :: nil => Some (Vfloat (Float.div f1 f2)) - | Osingleoffloat, Vfloat f1 :: nil => - Some (Vfloat (Float.singleoffloat f1)) - | Ointoffloat, Vfloat f1 :: nil => - Some (Vint (Float.intoffloat f1)) - | Ofloatofint, Vint n1 :: nil => - Some (Vfloat (Float.floatofint n1)) - | Ofloatofintu, Vint n1 :: nil => - Some (Vfloat (Float.floatofintu n1)) - | Ocmp c, Vint n1 :: Vint n2 :: nil => - Some (Val.of_bool(Int.cmp c n1 n2)) - | Ocmp c, Vptr b1 n1 :: Vptr b2 n2 :: nil => - if valid_pointer m b1 (Int.signed n1) - && valid_pointer m b2 (Int.signed n2) then - if eq_block b1 b2 then Some(Val.of_bool(Int.cmp c n1 n2)) else None - else - None - | Ocmp c, Vptr b1 n1 :: Vint n2 :: nil => eval_compare_null c n2 - | Ocmp c, Vint n1 :: Vptr b2 n2 :: nil => eval_compare_null c n1 - | Ocmpu c, Vint n1 :: Vint n2 :: nil => - Some (Val.of_bool(Int.cmpu c n1 n2)) - | Ocmpf c, Vfloat f1 :: Vfloat f2 :: nil => - Some (Val.of_bool (Float.cmp c f1 f2)) - | _, _ => None - end. - -(** Allocation of local variables at function entry. Each variable is - bound to the reference to a fresh block of the appropriate size. *) - -Inductive alloc_variables: env -> mem -> - list (ident * var_kind) -> - env -> mem -> list block -> Prop := - | alloc_variables_nil: - forall e m, - alloc_variables e m nil e m nil - | alloc_variables_cons: - forall e m id lv vars m1 b1 m2 e2 lb, - Mem.alloc m 0 (sizeof lv) = (m1, b1) -> - alloc_variables (PTree.set id (b1, lv) e) m1 vars e2 m2 lb -> - alloc_variables e m ((id, lv) :: vars) e2 m2 (b1 :: lb). - -(** Initialization of local variables that are parameters. The value - of the corresponding argument is stored into the memory block - bound to the parameter. *) - -Inductive bind_parameters: env -> - mem -> list (ident * memory_chunk) -> list val -> - mem -> Prop := - | bind_parameters_nil: - forall e m, - bind_parameters e m nil nil m - | bind_parameters_cons: - forall e m id chunk params v1 vl b m1 m2, - PTree.get id e = Some (b, Vscalar chunk) -> - Mem.store chunk m b 0 v1 = Some m1 -> - bind_parameters e m1 params vl m2 -> - bind_parameters e m ((id, chunk) :: params) (v1 :: vl) m2. - -Section RELSEM. - -Variable prg: program. -Let ge := Genv.globalenv (program_of_program prg). - -(* Evaluation of the address of a variable: - [eval_var_addr prg ge e id b] states that variable [id] - in environment [e] evaluates to block [b]. *) - -Inductive eval_var_addr: env -> ident -> block -> Prop := - | eval_var_addr_local: - forall e id b vi, - PTree.get id e = Some (b, vi) -> - eval_var_addr e id b - | eval_var_addr_global: - forall e id b, - PTree.get id e = None -> - Genv.find_symbol ge id = Some b -> - eval_var_addr e id b. - -(* Evaluation of a reference to a scalar variable: - [eval_var_ref prg ge e id b chunk] states - that variable [id] in environment [e] evaluates to block [b] - and is associated with the memory chunk [chunk]. *) - -Inductive eval_var_ref: env -> ident -> block -> memory_chunk -> Prop := - | eval_var_ref_local: - forall e id b chunk, - PTree.get id e = Some (b, Vscalar chunk) -> - eval_var_ref e id b chunk - | eval_var_ref_global: - forall e id b chunk, - PTree.get id e = None -> - Genv.find_symbol ge id = Some b -> - PTree.get id (global_var_env prg) = Some (Vscalar chunk) -> - eval_var_ref e id b chunk. - -(** Evaluation of an expression: [eval_expr prg le e m a m' v] states - that expression [a], in initial memory state [m], evaluates to value - [v]. [m'] is the final memory state, respectively, reflecting - memory stores possibly performed by [a]. [e] and [le] are the - local environment and let environment respectively. *) - -Inductive eval_expr: - letenv -> env -> - mem -> expr -> trace -> mem -> val -> Prop := - | eval_Evar: - forall le e m id b chunk v, - eval_var_ref e id b chunk -> - Mem.load chunk m b 0 = Some v -> - eval_expr le e m (Evar id) E0 m v - | eval_Eaddrof: - forall le e m id b, - eval_var_addr e id b -> - eval_expr le e m (Eaddrof id) E0 m (Vptr b Int.zero) - | eval_Eop: - forall le e m op al t m1 vl v, - eval_exprlist le e m al t m1 vl -> - eval_operation op vl m1 = Some v -> - eval_expr le e m (Eop op al) t m1 v - | eval_Eload: - forall le e m chunk a t m1 v1 v, - eval_expr le e m a t m1 v1 -> - Mem.loadv chunk m1 v1 = Some v -> - eval_expr le e m (Eload chunk a) t m1 v - | eval_Ecall: - forall le e m sig a bl t1 m1 t2 m2 t3 m3 vf vargs vres f t, - eval_expr le e m a t1 m1 vf -> - eval_exprlist le e m1 bl t2 m2 vargs -> - Genv.find_funct ge vf = Some f -> - funsig f = sig -> - eval_funcall m2 f vargs t3 m3 vres -> - t = t1 ** t2 ** t3 -> - eval_expr le e m (Ecall sig a bl) t m3 vres - | eval_Econdition_true: - forall le e m a b c t1 m1 v1 t2 m2 v2 t, - eval_expr le e m a t1 m1 v1 -> - Val.is_true v1 -> - eval_expr le e m1 b t2 m2 v2 -> - t = t1 ** t2 -> - eval_expr le e m (Econdition a b c) t m2 v2 - | eval_Econdition_false: - forall le e m a b c t1 m1 v1 t2 m2 v2 t, - eval_expr le e m a t1 m1 v1 -> - Val.is_false v1 -> - eval_expr le e m1 c t2 m2 v2 -> - t = t1 ** t2 -> - eval_expr le e m (Econdition a b c) t m2 v2 - | eval_Elet: - forall le e m a b t1 m1 v1 t2 m2 v2 t, - eval_expr le e m a t1 m1 v1 -> - eval_expr (v1::le) e m1 b t2 m2 v2 -> - t = t1 ** t2 -> - eval_expr le e m (Elet a b) t m2 v2 - | eval_Eletvar: - forall le e m n v, - nth_error le n = Some v -> - eval_expr le e m (Eletvar n) E0 m v - | eval_Ealloc: - forall le e m a t m1 n m2 b, - eval_expr le e m a t m1 (Vint n) -> - Mem.alloc m1 0 (Int.signed n) = (m2, b) -> - eval_expr le e m (Ealloc a) t m2 (Vptr b Int.zero) - -(** Evaluation of a list of expressions: - [eval_exprlist prg le al m a m' vl] - states that the list [al] of expressions evaluate - to the list [vl] of values. - The other parameters are as in [eval_expr]. -*) - -with eval_exprlist: - letenv -> env -> - mem -> exprlist -> trace -> - mem -> list val -> Prop := - | eval_Enil: - forall le e m, - eval_exprlist le e m Enil E0 m nil - | eval_Econs: - forall le e m a bl t1 m1 v t2 m2 vl t, - eval_expr le e m a t1 m1 v -> - eval_exprlist le e m1 bl t2 m2 vl -> - t = t1 ** t2 -> - eval_exprlist le e m (Econs a bl) t m2 (v :: vl) - -(** Evaluation of a function invocation: [eval_funcall prg m f args m' res] - means that the function [f], applied to the arguments [args] in - memory state [m], returns the value [res] in modified memory state [m']. -*) -with eval_funcall: - mem -> fundef -> list val -> trace -> - mem -> val -> Prop := - | eval_funcall_internal: - forall m f vargs e m1 lb m2 t m3 out vres, - list_norepet (fn_params_names f ++ fn_vars_names f) -> - alloc_variables empty_env m (fn_variables f) e m1 lb -> - bind_parameters e m1 f.(fn_params) vargs m2 -> - exec_stmt e m2 f.(fn_body) t m3 out -> - outcome_result_value out f.(fn_sig).(sig_res) vres -> - eval_funcall m (Internal f) vargs t (Mem.free_list m3 lb) vres - | eval_funcall_external: - forall m ef vargs t vres, - event_match ef vargs t vres -> - eval_funcall m (External ef) vargs t m vres - -(** Execution of a statement: [exec_stmt prg e m s m' out] - means that statement [s] executes with outcome [out]. - The other parameters are as in [eval_expr]. *) - -with exec_stmt: - env -> - mem -> stmt -> trace -> - mem -> outcome -> Prop := - | exec_Sskip: - forall e m, - exec_stmt e m Sskip E0 m Out_normal - | exec_Sexpr: - forall e m a t m1 v, - eval_expr nil e m a t m1 v -> - exec_stmt e m (Sexpr a) t m1 Out_normal - | eval_Sassign: - forall e m id a t m1 b chunk v m2, - eval_expr nil e m a t m1 v -> - eval_var_ref e id b chunk -> - Mem.store chunk m1 b 0 v = Some m2 -> - exec_stmt e m (Sassign id a) t m2 Out_normal - | eval_Sstore: - forall e m chunk a b t1 m1 v1 t2 m2 v2 t3 m3, - eval_expr nil e m a t1 m1 v1 -> - eval_expr nil e m1 b t2 m2 v2 -> - Mem.storev chunk m2 v1 v2 = Some m3 -> - t3 = t1 ** t2 -> - exec_stmt e m (Sstore chunk a b) t3 m3 Out_normal - | exec_Sseq_continue: - forall e m s1 s2 t1 t2 m1 m2 t out, - exec_stmt e m s1 t1 m1 Out_normal -> - exec_stmt e m1 s2 t2 m2 out -> - t = t1 ** t2 -> - exec_stmt e m (Sseq s1 s2) t m2 out - | exec_Sseq_stop: - forall e m s1 s2 t1 m1 out, - exec_stmt e m s1 t1 m1 out -> - out <> Out_normal -> - exec_stmt e m (Sseq s1 s2) t1 m1 out - | exec_Sifthenelse_true: - forall e m a sl1 sl2 t1 m1 v1 t2 m2 out t, - eval_expr nil e m a t1 m1 v1 -> - Val.is_true v1 -> - exec_stmt e m1 sl1 t2 m2 out -> - t = t1 ** t2 -> - exec_stmt e m (Sifthenelse a sl1 sl2) t m2 out - | exec_Sifthenelse_false: - forall e m a sl1 sl2 t1 m1 v1 t2 m2 out t, - eval_expr nil e m a t1 m1 v1 -> - Val.is_false v1 -> - exec_stmt e m1 sl2 t2 m2 out -> - t = t1 ** t2 -> - exec_stmt e m (Sifthenelse a sl1 sl2) t m2 out - | exec_Sloop_loop: - forall e m sl t1 m1 t2 m2 out t, - exec_stmt e m sl t1 m1 Out_normal -> - exec_stmt e m1 (Sloop sl) t2 m2 out -> - t = t1 ** t2 -> - exec_stmt e m (Sloop sl) t m2 out - | exec_Sloop_stop: - forall e m sl t1 m1 out, - exec_stmt e m sl t1 m1 out -> - out <> Out_normal -> - exec_stmt e m (Sloop sl) t1 m1 out - | exec_Sblock: - forall e m sl t1 m1 out, - exec_stmt e m sl t1 m1 out -> - exec_stmt e m (Sblock sl) t1 m1 (outcome_block out) - | exec_Sexit: - forall e m n, - exec_stmt e m (Sexit n) E0 m (Out_exit n) - | exec_Sswitch: - forall e m a cases default t1 m1 n, - eval_expr nil e m a t1 m1 (Vint n) -> - exec_stmt e m (Sswitch a cases default) - t1 m1 (Out_exit (switch_target n default cases)) - | exec_Sreturn_none: - forall e m, - exec_stmt e m (Sreturn None) E0 m (Out_return None) - | exec_Sreturn_some: - forall e m a t1 m1 v, - eval_expr nil e m a t1 m1 v -> - exec_stmt e m (Sreturn (Some a)) t1 m1 (Out_return (Some v)). - -Scheme eval_expr_ind4 := Minimality for eval_expr Sort Prop - with eval_exprlist_ind4 := Minimality for eval_exprlist Sort Prop - with eval_funcall_ind4 := Minimality for eval_funcall Sort Prop - with exec_stmt_ind4 := Minimality for exec_stmt Sort Prop. - -End RELSEM. - -(** Execution of a whole program: [exec_program p r] - holds if the application of [p]'s main function to no arguments - in the initial memory state for [p] eventually returns value [r]. *) - -Definition exec_program (p: program) (t: trace) (r: val) : Prop := - let ge := Genv.globalenv (program_of_program p) in - let m0 := Genv.init_mem (program_of_program p) in - exists b, exists f, exists m, - Genv.find_symbol ge p.(prog_main) = Some b /\ - Genv.find_funct_ptr ge b = Some f /\ - funsig f = mksignature nil (Some Tint) /\ - eval_funcall p m0 f nil t m r. diff --git a/backend/Events.v b/backend/Events.v deleted file mode 100644 index a0559fd0..00000000 --- a/backend/Events.v +++ /dev/null @@ -1,103 +0,0 @@ -(** Representation of (traces of) observable events. *) - -Require Import Coqlib. -Require Import AST. -Require Import Integers. -Require Import Floats. -Require Import Values. - -Inductive eventval: Set := - | EVint: int -> eventval - | EVfloat: float -> eventval. - -Parameter trace: Set. -Parameter E0: trace. -Parameter Eextcall: ident -> list eventval -> eventval -> trace. -Parameter Eapp: trace -> trace -> trace. - -Infix "**" := Eapp (at level 60, right associativity). - -Axiom E0_left: forall t, E0 ** t = t. -Axiom E0_right: forall t, t ** E0 = t. -Axiom Eapp_assoc: forall t1 t2 t3, (t1 ** t2) ** t3 = t1 ** (t2 ** t3). - -Hint Rewrite E0_left E0_right Eapp_assoc: trace_rewrite. - -Ltac substTraceHyp := - match goal with - | [ H: (@eq trace ?x ?y) |- _ ] => - subst x || clear H - end. - -Ltac decomposeTraceEq := - match goal with - | [ |- (_ ** _) = (_ ** _) ] => - apply (f_equal2 Eapp); auto; decomposeTraceEq - | _ => - auto - end. - -Ltac traceEq := - repeat substTraceHyp; autorewrite with trace_rewrite; decomposeTraceEq. - -Inductive eventval_match: eventval -> typ -> val -> Prop := - | ev_match_int: - forall i, eventval_match (EVint i) Tint (Vint i) - | ev_match_float: - forall f, eventval_match (EVfloat f) Tfloat (Vfloat f). - -Inductive eventval_list_match: list eventval -> list typ -> list val -> Prop := - | evl_match_nil: - eventval_list_match nil nil nil - | evl_match_cons: - forall ev1 evl ty1 tyl v1 vl, - eventval_match ev1 ty1 v1 -> - eventval_list_match evl tyl vl -> - eventval_list_match (ev1::evl) (ty1::tyl) (v1::vl). - -Inductive event_match: - external_function -> list val -> trace -> val -> Prop := - event_match_intro: - forall ef vargs vres eargs eres, - eventval_list_match eargs (sig_args ef.(ef_sig)) vargs -> - eventval_match eres (proj_sig_res ef.(ef_sig)) vres -> - event_match ef vargs (Eextcall ef.(ef_id) eargs eres) vres. - -Require Import Mem. - -Section EVENT_MATCH_INJECT. - -Variable f: meminj. - -Remark eventval_match_inject: - forall ev ty v1, eventval_match ev ty v1 -> - forall v2, val_inject f v1 v2 -> - eventval_match ev ty v2. -Proof. - induction 1; intros; inversion H; constructor. -Qed. - -Remark eventval_list_match_inject: - forall evl tyl vl1, eventval_list_match evl tyl vl1 -> - forall vl2, val_list_inject f vl1 vl2 -> - eventval_list_match evl tyl vl2. -Proof. - induction 1; intros. - inversion H; constructor. - inversion H1; constructor. - eapply eventval_match_inject; eauto. - eauto. -Qed. - -Lemma event_match_inject: - forall ef args1 t res args2, - event_match ef args1 t res -> - val_list_inject f args1 args2 -> - event_match ef args2 t res /\ val_inject f res res. -Proof. - intros. inversion H; subst. - split. constructor. eapply eventval_list_match_inject; eauto. auto. - inversion H2; constructor. -Qed. - -End EVENT_MATCH_INJECT. diff --git a/backend/Globalenvs.v b/backend/Globalenvs.v deleted file mode 100644 index 036fd8f6..00000000 --- a/backend/Globalenvs.v +++ /dev/null @@ -1,643 +0,0 @@ -(** Global environments are a component of the dynamic semantics of - all languages involved in the compiler. A global environment - maps symbol names (names of functions and of global variables) - to the corresponding memory addresses. It also maps memory addresses - of functions to the corresponding function descriptions. - - Global environments, along with the initial memory state at the beginning - of program execution, are built from the program of interest, as follows: -- A distinct memory address is assigned to each function of the program. - These function addresses use negative numbers to distinguish them from - addresses of memory blocks. The associations of function name to function - address and function address to function description are recorded in - the global environment. -- For each global variable, a memory block is allocated and associated to - the name of the variable. - - These operations reflect (at a high level of abstraction) what takes - place during program linking and program loading in a real operating - system. *) - -Require Import Coqlib. -Require Import Maps. -Require Import AST. -Require Import Integers. -Require Import Values. -Require Import Mem. - -Set Implicit Arguments. - -Module Type GENV. - -(** ** Types and operations *) - - Variable t: Set -> Set. - (** The type of global environments. The parameter [F] is the type - of function descriptions. *) - - Variable globalenv: forall (F: Set), program F -> t F. - (** Return the global environment for the given program. *) - - Variable init_mem: forall (F: Set), program F -> mem. - (** Return the initial memory state for the given program. *) - - Variable find_funct_ptr: forall (F: Set), t F -> block -> option F. - (** Return the function description associated with the given address, - if any. *) - - Variable find_funct: forall (F: Set), t F -> val -> option F. - (** Same as [find_funct_ptr] but the function address is given as - a value, which must be a pointer with offset 0. *) - - Variable find_symbol: forall (F: Set), t F -> ident -> option block. - (** Return the address of the given global symbol, if any. *) - -(** ** Properties of the operations. *) - - Hypothesis find_funct_inv: - forall (F: Set) (ge: t F) (v: val) (f: F), - find_funct ge v = Some f -> exists b, v = Vptr b Int.zero. - Hypothesis find_funct_find_funct_ptr: - forall (F: Set) (ge: t F) (b: block), - find_funct ge (Vptr b Int.zero) = find_funct_ptr ge b. - Hypothesis find_funct_ptr_prop: - forall (F: Set) (P: F -> Prop) (p: program F) (b: block) (f: F), - (forall id f, In (id, f) (prog_funct p) -> P f) -> - find_funct_ptr (globalenv p) b = Some f -> - P f. - Hypothesis find_funct_prop: - forall (F: Set) (P: F -> Prop) (p: program F) (v: val) (f: F), - (forall id f, In (id, f) (prog_funct p) -> P f) -> - find_funct (globalenv p) v = Some f -> - P f. - Hypothesis find_funct_ptr_symbol_inversion: - forall (F: Set) (p: program F) (id: ident) (b: block) (f: F), - find_symbol (globalenv p) id = Some b -> - find_funct_ptr (globalenv p) b = Some f -> - In (id, f) p.(prog_funct). - - Hypothesis initmem_nullptr: - forall (F: Set) (p: program F), - let m := init_mem p in - valid_block m nullptr /\ - m.(blocks) nullptr = empty_block 0 0. - Hypothesis initmem_block_init: - forall (F: Set) (p: program F) (b: block), - exists id, (init_mem p).(blocks) b = block_init_data id. - Hypothesis find_funct_ptr_inv: - forall (F: Set) (p: program F) (b: block) (f: F), - find_funct_ptr (globalenv p) b = Some f -> b < 0. - Hypothesis find_symbol_inv: - forall (F: Set) (p: program F) (id: ident) (b: block), - find_symbol (globalenv p) id = Some b -> b < nextblock (init_mem p). - -(** Commutation properties between program transformations - and operations over global environments. *) - - Hypothesis find_funct_ptr_transf: - forall (A B: Set) (transf: A -> B) (p: program A) (b: block) (f: A), - find_funct_ptr (globalenv p) b = Some f -> - find_funct_ptr (globalenv (transform_program transf p)) b = Some (transf f). - Hypothesis find_funct_transf: - forall (A B: Set) (transf: A -> B) (p: program A) (v: val) (f: A), - find_funct (globalenv p) v = Some f -> - find_funct (globalenv (transform_program transf p)) v = Some (transf f). - Hypothesis find_symbol_transf: - forall (A B: Set) (transf: A -> B) (p: program A) (s: ident), - find_symbol (globalenv (transform_program transf p)) s = - find_symbol (globalenv p) s. - Hypothesis init_mem_transf: - forall (A B: Set) (transf: A -> B) (p: program A), - init_mem (transform_program transf p) = init_mem p. - -(** Commutation properties between partial program transformations - and operations over global environments. *) - - Hypothesis find_funct_ptr_transf_partial: - forall (A B: Set) (transf: A -> option B) - (p: program A) (p': program B), - transform_partial_program transf p = Some p' -> - forall (b: block) (f: A), - find_funct_ptr (globalenv p) b = Some f -> - find_funct_ptr (globalenv p') b = transf f /\ transf f <> None. - Hypothesis find_funct_transf_partial: - forall (A B: Set) (transf: A -> option B) - (p: program A) (p': program B), - transform_partial_program transf p = Some p' -> - forall (v: val) (f: A), - find_funct (globalenv p) v = Some f -> - find_funct (globalenv p') v = transf f /\ transf f <> None. - Hypothesis find_symbol_transf_partial: - forall (A B: Set) (transf: A -> option B) - (p: program A) (p': program B), - transform_partial_program transf p = Some p' -> - forall (s: ident), - find_symbol (globalenv p') s = find_symbol (globalenv p) s. - Hypothesis init_mem_transf_partial: - forall (A B: Set) (transf: A -> option B) - (p: program A) (p': program B), - transform_partial_program transf p = Some p' -> - init_mem p' = init_mem p. -End GENV. - -(** The rest of this library is a straightforward implementation of - the module signature above. *) - -Module Genv: GENV. - -Section GENV. - -Variable funct: Set. (* The type of functions *) - -Record genv : Set := mkgenv { - functions: ZMap.t (option funct); (* mapping function ptr -> function *) - nextfunction: Z; - symbols: PTree.t block (* mapping symbol -> block *) -}. - -Definition t := genv. - -Definition add_funct (name_fun: (ident * funct)) (g: genv) : genv := - let b := g.(nextfunction) in - mkgenv (ZMap.set b (Some (snd name_fun)) g.(functions)) - (Zpred b) - (PTree.set (fst name_fun) b g.(symbols)). - -Definition add_symbol (name: ident) (b: block) (g: genv) : genv := - mkgenv g.(functions) - g.(nextfunction) - (PTree.set name b g.(symbols)). - -Definition find_funct_ptr (g: genv) (b: block) : option funct := - ZMap.get b g.(functions). - -Definition find_funct (g: genv) (v: val) : option funct := - match v with - | Vptr b ofs => - if Int.eq ofs Int.zero then find_funct_ptr g b else None - | _ => - None - end. - -Definition find_symbol (g: genv) (symb: ident) : option block := - PTree.get symb g.(symbols). - -Lemma find_funct_inv: - forall (ge: t) (v: val) (f: funct), - find_funct ge v = Some f -> exists b, v = Vptr b Int.zero. -Proof. - intros until f. unfold find_funct. destruct v; try (intros; discriminate). - generalize (Int.eq_spec i Int.zero). case (Int.eq i Int.zero); intros. - exists b. congruence. - discriminate. -Qed. - -Lemma find_funct_find_funct_ptr: - forall (ge: t) (b: block), - find_funct ge (Vptr b Int.zero) = find_funct_ptr ge b. -Proof. - intros. simpl. - generalize (Int.eq_spec Int.zero Int.zero). - case (Int.eq Int.zero Int.zero); intros. - auto. tauto. -Qed. - -(* Construct environment and initial memory store *) - -Definition empty : genv := - mkgenv (ZMap.init None) (-1) (PTree.empty block). - -Definition add_functs (init: genv) (fns: list (ident * funct)) : genv := - List.fold_right add_funct init fns. - -Definition add_globals - (init: genv * mem) (vars: list (ident * list init_data)) : genv * mem := - List.fold_right - (fun (id_init: ident * list init_data) (g_st: genv * mem) => - let (id, init) := id_init in - let (g, st) := g_st in - let (st', b) := Mem.alloc_init_data st init in - (add_symbol id b g, st')) - init vars. - -Definition globalenv_initmem (p: program funct) : (genv * mem) := - add_globals - (add_functs empty p.(prog_funct), Mem.empty) - p.(prog_vars). - -Definition globalenv (p: program funct) : genv := - fst (globalenv_initmem p). -Definition init_mem (p: program funct) : mem := - snd (globalenv_initmem p). - -Lemma functions_globalenv: - forall (p: program funct), - functions (globalenv p) = functions (add_functs empty p.(prog_funct)). -Proof. - assert (forall init vars, - functions (fst (add_globals init vars)) = functions (fst init)). - induction vars; simpl. - reflexivity. - destruct a. destruct (add_globals init vars). - simpl. exact IHvars. - - unfold add_globals; simpl. - intros. unfold globalenv; unfold globalenv_initmem. - rewrite H. reflexivity. -Qed. - -Lemma initmem_nullptr: - forall (p: program funct), - let m := init_mem p in - valid_block m nullptr /\ - m.(blocks) nullptr = mkblock 0 0 (fun y => Undef) (undef_undef_outside 0 0). -Proof. - assert - (forall init, - let m1 := snd init in - 0 < m1.(nextblock) -> - m1.(blocks) nullptr = mkblock 0 0 (fun y => Undef) (undef_undef_outside 0 0) -> - forall vars, - let m2 := snd (add_globals init vars) in - 0 < m2.(nextblock) /\ - m2.(blocks) nullptr = mkblock 0 0 (fun y => Undef) (undef_undef_outside 0 0)). - induction vars; simpl; intros. - tauto. - destruct a. - caseEq (add_globals init vars). intros g m2 EQ. - rewrite EQ in IHvars. simpl in IHvars. elim IHvars; intros. - simpl. split. omega. - rewrite update_o. auto. apply sym_not_equal. apply Zlt_not_eq. exact H1. - - intro. unfold init_mem. unfold globalenv_initmem. - unfold valid_block. apply H. simpl. omega. reflexivity. -Qed. - -Lemma initmem_block_init: - forall (p: program funct) (b: block), - exists id, (init_mem p).(blocks) b = block_init_data id. -Proof. - assert (forall g0 vars g1 m b, - add_globals (g0, Mem.empty) vars = (g1, m) -> - exists id, m.(blocks) b = block_init_data id). - induction vars; simpl. - intros. inversion H. unfold Mem.empty; simpl. - exists (@nil init_data). symmetry. apply Mem.block_init_data_empty. - destruct a. caseEq (add_globals (g0, Mem.empty) vars). intros g1 m1 EQ. - intros g m b EQ1. injection EQ1; intros EQ2 EQ3; clear EQ1. - rewrite <- EQ2; simpl. unfold update. - case (zeq b (nextblock m1)); intro. - exists l; auto. - eauto. - intros. caseEq (globalenv_initmem p). - intros g m EQ. unfold init_mem; rewrite EQ; simpl. - unfold globalenv_initmem in EQ. eauto. -Qed. - -Remark nextfunction_add_functs_neg: - forall fns, nextfunction (add_functs empty fns) < 0. -Proof. - induction fns; simpl; intros. omega. unfold Zpred. omega. -Qed. - -Theorem find_funct_ptr_inv: - forall (p: program funct) (b: block) (f: funct), - find_funct_ptr (globalenv p) b = Some f -> b < 0. -Proof. - intros until f. - assert (forall fns, ZMap.get b (functions (add_functs empty fns)) = Some f -> b < 0). - induction fns; simpl. - rewrite ZMap.gi. congruence. - rewrite ZMap.gsspec. case (ZIndexed.eq b (nextfunction (add_functs empty fns))); intro. - intro. rewrite e. apply nextfunction_add_functs_neg. - auto. - unfold find_funct_ptr. rewrite functions_globalenv. - intros. eauto. -Qed. - -Theorem find_symbol_inv: - forall (p: program funct) (id: ident) (b: block), - find_symbol (globalenv p) id = Some b -> b < nextblock (init_mem p). -Proof. - assert (forall fns s b, - (symbols (add_functs empty fns)) ! s = Some b -> b < 0). - induction fns; simpl; intros until b. - rewrite PTree.gempty. congruence. - rewrite PTree.gsspec. destruct a; simpl. case (peq s i); intro. - intro EQ; inversion EQ. apply nextfunction_add_functs_neg. - eauto. - assert (forall fns vars g m s b, - add_globals (add_functs empty fns, Mem.empty) vars = (g, m) -> - (symbols g)!s = Some b -> - b < nextblock m). - induction vars; simpl; intros until b. - intros. inversion H0. subst g m. simpl. - generalize (H fns s b H1). omega. - destruct a. caseEq (add_globals (add_functs empty fns, Mem.empty) vars). - intros g1 m1 ADG EQ. inversion EQ; subst g m; clear EQ. - unfold add_symbol; simpl. rewrite PTree.gsspec. case (peq s i); intro. - intro EQ; inversion EQ. omega. - intro. generalize (IHvars _ _ _ _ ADG H0). omega. - intros until b. unfold find_symbol, globalenv, init_mem, globalenv_initmem; simpl. - caseEq (add_globals (add_functs empty (prog_funct p), Mem.empty) - (prog_vars p)); intros g m EQ. - simpl; intros. eauto. -Qed. - -End GENV. - -(* Invariants on functions *) -Lemma find_funct_ptr_prop: - forall (F: Set) (P: F -> Prop) (p: program F) (b: block) (f: F), - (forall id f, In (id, f) (prog_funct p) -> P f) -> - find_funct_ptr (globalenv p) b = Some f -> - P f. -Proof. - intros until f. - unfold find_funct_ptr. rewrite functions_globalenv. - generalize (prog_funct p). induction l; simpl. - rewrite ZMap.gi. intros; discriminate. - rewrite ZMap.gsspec. - case (ZIndexed.eq b (nextfunction (add_functs (empty F) l))); intros. - apply H with (fst a). left. destruct a. simpl in *. congruence. - apply IHl. intros. apply H with id. right. auto. auto. -Qed. - -Lemma find_funct_prop: - forall (F: Set) (P: F -> Prop) (p: program F) (v: val) (f: F), - (forall id f, In (id, f) (prog_funct p) -> P f) -> - find_funct (globalenv p) v = Some f -> - P f. -Proof. - intros until f. unfold find_funct. - destruct v; try (intros; discriminate). - case (Int.eq i Int.zero); [idtac | intros; discriminate]. - intros. eapply find_funct_ptr_prop; eauto. -Qed. - -Lemma find_funct_ptr_symbol_inversion: - forall (F: Set) (p: program F) (id: ident) (b: block) (f: F), - find_symbol (globalenv p) id = Some b -> - find_funct_ptr (globalenv p) b = Some f -> - In (id, f) p.(prog_funct). -Proof. - intros until f. - assert (forall fns, - let g := add_functs (empty F) fns in - PTree.get id g.(symbols) = Some b -> - b > g.(nextfunction)). - induction fns; simpl. - rewrite PTree.gempty. congruence. - rewrite PTree.gsspec. case (peq id (fst a)); intro. - intro EQ. inversion EQ. unfold Zpred. omega. - intros. generalize (IHfns H). unfold Zpred; omega. - assert (forall fns, - let g := add_functs (empty F) fns in - PTree.get id g.(symbols) = Some b -> - ZMap.get b g.(functions) = Some f -> - In (id, f) fns). - induction fns; simpl. - rewrite ZMap.gi. congruence. - set (g := add_functs (empty F) fns). - rewrite PTree.gsspec. rewrite ZMap.gsspec. - case (peq id (fst a)); intro. - intro EQ. inversion EQ. unfold ZIndexed.eq. rewrite zeq_true. - intro EQ2. left. destruct a. simpl in *. congruence. - intro. unfold ZIndexed.eq. rewrite zeq_false. intro. eauto. - generalize (H _ H0). fold g. unfold block. omega. - assert (forall g0 m0, b < 0 -> - forall vars g m, - @add_globals F (g0, m0) vars = (g, m) -> - PTree.get id g.(symbols) = Some b -> - PTree.get id g0.(symbols) = Some b). - induction vars; simpl. - intros. inversion H2. auto. - destruct a. caseEq (add_globals (g0, m0) vars). - intros g1 m1 EQ g m EQ1. injection EQ1; simpl; clear EQ1. - unfold add_symbol; intros A B. rewrite <- B. simpl. - rewrite PTree.gsspec. case (peq id i); intros. - assert (b > 0). injection H2; intros. rewrite <- H3. apply nextblock_pos. - omegaContradiction. - eauto. - intros. - generalize (find_funct_ptr_inv _ _ H3). intro. - pose (g := add_functs (empty F) (prog_funct p)). - apply H0. - apply H1 with Mem.empty (prog_vars p) (globalenv p) (init_mem p). - auto. unfold globalenv, init_mem. rewrite <- surjective_pairing. - reflexivity. assumption. rewrite <- functions_globalenv. assumption. -Qed. - -(* Global environments and program transformations. *) - -Section TRANSF_PROGRAM_PARTIAL. - -Variable A B: Set. -Variable transf: A -> option B. -Variable p: program A. -Variable p': program B. -Hypothesis transf_OK: transform_partial_program transf p = Some p'. - -Lemma add_functs_transf: - forall (fns: list (ident * A)) (tfns: list (ident * B)), - transf_partial_program transf fns = Some tfns -> - let r := add_functs (empty A) fns in - let tr := add_functs (empty B) tfns in - nextfunction tr = nextfunction r /\ - symbols tr = symbols r /\ - forall (b: block) (f: A), - ZMap.get b (functions r) = Some f -> - ZMap.get b (functions tr) = transf f /\ transf f <> None. -Proof. - induction fns; simpl. - - intros; injection H; intro; subst tfns. - simpl. split. reflexivity. split. reflexivity. - intros b f; repeat (rewrite ZMap.gi). intros; discriminate. - - intro tfns. destruct a. caseEq (transf a). intros a' TA. - caseEq (transf_partial_program transf fns). intros l TPP EQ. - injection EQ; intro; subst tfns. - clear EQ. simpl. - generalize (IHfns l TPP). - intros [HR1 [HR2 HR3]]. - rewrite HR1. rewrite HR2. - split. reflexivity. - split. reflexivity. - intros b f. - case (zeq b (nextfunction (add_functs (empty A) fns))); intro. - subst b. repeat (rewrite ZMap.gss). - intro EQ; injection EQ; intro; subst f; clear EQ. - rewrite TA. split. auto. discriminate. - repeat (rewrite ZMap.gso; auto). - - intros; discriminate. - intros; discriminate. -Qed. - -Lemma mem_add_globals_transf: - forall (g1: genv A) (g2: genv B) (m: mem) (vars: list (ident * list init_data)), - snd (add_globals (g1, m) vars) = snd (add_globals (g2, m) vars). -Proof. - induction vars; simpl. - reflexivity. - destruct a. destruct (add_globals (g1, m) vars). - destruct (add_globals (g2, m) vars). - simpl in IHvars. subst m1. reflexivity. -Qed. - -Lemma symbols_add_globals_transf: - forall (g1: genv A) (g2: genv B) (m: mem), - symbols g1 = symbols g2 -> - forall (vars: list (ident * list init_data)), - symbols (fst (add_globals (g1, m) vars)) = - symbols (fst (add_globals (g2, m) vars)). -Proof. - induction vars; simpl. - assumption. - generalize (mem_add_globals_transf g1 g2 m vars); intro. - destruct a. destruct (add_globals (g1, m) vars). - destruct (add_globals (g2, m) vars). - simpl. simpl in IHvars. simpl in H0. - rewrite H0; rewrite IHvars. reflexivity. -Qed. - -Lemma prog_funct_transf_OK: - transf_partial_program transf p.(prog_funct) = Some p'.(prog_funct). -Proof. - generalize transf_OK; unfold transform_partial_program. - case (transf_partial_program transf (prog_funct p)); simpl; intros. - injection transf_OK0; intros; subst p'. reflexivity. - discriminate. -Qed. - -Theorem find_funct_ptr_transf_partial: - forall (b: block) (f: A), - find_funct_ptr (globalenv p) b = Some f -> - find_funct_ptr (globalenv p') b = transf f /\ transf f <> None. -Proof. - intros until f. - generalize (add_functs_transf p.(prog_funct) prog_funct_transf_OK). - intros [X [Y Z]]. - unfold find_funct_ptr. - repeat (rewrite functions_globalenv). - apply Z. -Qed. - -Theorem find_funct_transf_partial: - forall (v: val) (f: A), - find_funct (globalenv p) v = Some f -> - find_funct (globalenv p') v = transf f /\ transf f <> None. -Proof. - intros until f. unfold find_funct. - case v; try (intros; discriminate). - intros b ofs. - case (Int.eq ofs Int.zero); try (intros; discriminate). - apply find_funct_ptr_transf_partial. -Qed. - -Lemma symbols_init_transf: - symbols (globalenv p') = symbols (globalenv p). -Proof. - unfold globalenv. unfold globalenv_initmem. - generalize (add_functs_transf p.(prog_funct) prog_funct_transf_OK). - intros [X [Y Z]]. - generalize transf_OK. - unfold transform_partial_program. - case (transf_partial_program transf (prog_funct p)). - intros. injection transf_OK0; intro; subst p'; simpl. - symmetry. apply symbols_add_globals_transf. - symmetry. exact Y. - intros; discriminate. -Qed. - -Theorem find_symbol_transf_partial: - forall (s: ident), - find_symbol (globalenv p') s = find_symbol (globalenv p) s. -Proof. - intros. unfold find_symbol. - rewrite symbols_init_transf. auto. -Qed. - -Theorem init_mem_transf_partial: - init_mem p' = init_mem p. -Proof. - unfold init_mem. unfold globalenv_initmem. - generalize transf_OK. - unfold transform_partial_program. - case (transf_partial_program transf (prog_funct p)). - intros. injection transf_OK0; intro; subst p'; simpl. - symmetry. apply mem_add_globals_transf. - intros; discriminate. -Qed. - -End TRANSF_PROGRAM_PARTIAL. - -Section TRANSF_PROGRAM. - -Variable A B: Set. -Variable transf: A -> B. -Variable p: program A. -Let tp := transform_program transf p. - -Definition transf_partial (x: A) : option B := Some (transf x). - -Lemma transf_program_transf_partial_program: - forall (fns: list (ident * A)), - transf_partial_program transf_partial fns = - Some (transf_program transf fns). -Proof. - induction fns; simpl. - reflexivity. - destruct a. rewrite IHfns. reflexivity. -Qed. - -Lemma transform_program_transform_partial_program: - transform_partial_program transf_partial p = Some tp. -Proof. - unfold tp. unfold transform_partial_program, transform_program. - rewrite transf_program_transf_partial_program. - reflexivity. -Qed. - -Theorem find_funct_ptr_transf: - forall (b: block) (f: A), - find_funct_ptr (globalenv p) b = Some f -> - find_funct_ptr (globalenv tp) b = Some (transf f). -Proof. - intros. - generalize (find_funct_ptr_transf_partial transf_partial p - transform_program_transform_partial_program). - intros. elim (H0 b f H). intros. exact H1. -Qed. - -Theorem find_funct_transf: - forall (v: val) (f: A), - find_funct (globalenv p) v = Some f -> - find_funct (globalenv tp) v = Some (transf f). -Proof. - intros. - generalize (find_funct_transf_partial transf_partial p - transform_program_transform_partial_program). - intros. elim (H0 v f H). intros. exact H1. -Qed. - -Theorem find_symbol_transf: - forall (s: ident), - find_symbol (globalenv tp) s = find_symbol (globalenv p) s. -Proof. - intros. - apply find_symbol_transf_partial with transf_partial. - apply transform_program_transform_partial_program. -Qed. - -Theorem init_mem_transf: - init_mem tp = init_mem p. -Proof. - apply init_mem_transf_partial with transf_partial. - apply transform_program_transform_partial_program. -Qed. - -End TRANSF_PROGRAM. - -End Genv. diff --git a/backend/Main.v b/backend/Main.v deleted file mode 100644 index 95dc4e6c..00000000 --- a/backend/Main.v +++ /dev/null @@ -1,311 +0,0 @@ -(** The compiler back-end and its proof of semantic preservation *) - -(** Libraries. *) -Require Import Coqlib. -Require Import Maps. -Require Import AST. -Require Import Values. -(** Languages (syntax and semantics). *) -Require Csharpminor. -Require Cminor. -Require RTL. -Require LTL. -Require Linear. -Require Mach. -Require PPC. -(** Translation passes. *) -Require Cminorgen. -Require RTLgen. -Require Constprop. -Require CSE. -Require Allocation. -Require Tunneling. -Require Linearize. -Require Stacking. -Require PPCgen. -(** Type systems. *) -Require RTLtyping. -Require LTLtyping. -Require Lineartyping. -Require Machtyping. -(** Proofs of semantic preservation and typing preservation. *) -Require Cminorgenproof. -Require RTLgenproof. -Require Constpropproof. -Require CSEproof. -Require Allocproof. -Require Alloctyping. -Require Tunnelingproof. -Require Tunnelingtyping. -Require Linearizeproof. -Require Linearizetyping. -Require Stackingproof. -Require Stackingtyping. -Require Machabstr2mach. -Require PPCgenproof. - -(** * Composing the translation passes *) - -(** We first define useful monadic composition operators, - along with funny (but convenient) notations. *) - -Definition apply_total (A B: Set) (x: option A) (f: A -> B) : option B := - match x with None => None | Some x1 => Some (f x1) end. - -Definition apply_partial (A B: Set) - (x: option A) (f: A -> option B) : option B := - match x with None => None | Some x1 => f x1 end. - -Notation "a @@@ b" := - (apply_partial _ _ a b) (at level 50, left associativity). -Notation "a @@ b" := - (apply_total _ _ a b) (at level 50, left associativity). - -(** We define two translation functions for whole programs: one starting with - a Csharpminor program, the other with a Cminor program. Both - translations produce PPC programs ready for pretty-printing and - assembling. Some front-ends will prefer to generate Csharpminor - (e.g. a C front-end) while some others (e.g. an ML compiler) might - find it more convenient to generate Cminor directly, bypassing - Csharpminor. - - There are two ways to compose the compiler passes. The first translates - every function from the Cminor program from Cminor to RTL, then to LTL, etc, - all the way to PPC, and iterates this transformation for every function. - The second translates the whole Cminor program to a RTL program, then to - an LTL program, etc. We follow the first approach because it has lower - memory requirements. - - The translation of a Cminor function to a PPC function is as follows. *) - -Definition transf_cminor_fundef (f: Cminor.fundef) : option PPC.fundef := - Some f - @@@ RTLgen.transl_fundef - @@ Constprop.transf_fundef - @@ CSE.transf_fundef - @@@ Allocation.transf_fundef - @@ Tunneling.tunnel_fundef - @@ Linearize.transf_fundef - @@@ Stacking.transf_fundef - @@@ PPCgen.transf_fundef. - -(** And here is the translation for Csharpminor functions. *) - -Definition transf_csharpminor_fundef - (gce: Cminorgen.compilenv) (f: Csharpminor.fundef) : option PPC.fundef := - Some f - @@@ Cminorgen.transl_fundef gce - @@@ transf_cminor_fundef. - -(** The corresponding translations for whole program follow. *) - -Definition transf_cminor_program (p: Cminor.program) : option PPC.program := - transform_partial_program transf_cminor_fundef p. - -Definition transf_csharpminor_program (p: Csharpminor.program) : option PPC.program := - let gce := Cminorgen.build_global_compilenv p in - transform_partial_program - (transf_csharpminor_fundef gce) - (Csharpminor.program_of_program p). - -(** * Equivalence with whole program transformations *) - -(** To prove semantic preservation for the whole compiler, it is easier to reason - over the second way to compose the compiler pass: the one that translate - whole programs through each compiler pass. We now define this second translation - and prove that it produces the same PPC programs as the first translation. *) - -Definition transf_cminor_program2 (p: Cminor.program) : option PPC.program := - Some p - @@@ RTLgen.transl_program - @@ Constprop.transf_program - @@ CSE.transf_program - @@@ Allocation.transf_program - @@ Tunneling.tunnel_program - @@ Linearize.transf_program - @@@ Stacking.transf_program - @@@ PPCgen.transf_program. - -Definition transf_csharpminor_program2 (p: Csharpminor.program) : option PPC.program := - Some p - @@@ Cminorgen.transl_program - @@@ transf_cminor_program2. - -Lemma transf_partial_program_compose: - forall (A B C: Set) - (f1: A -> option B) (f2: B -> option C) - (p: list (ident * A)), - transf_partial_program f1 p @@@ transf_partial_program f2 = - transf_partial_program (fun f => f1 f @@@ f2) p. -Proof. - induction p. simpl. auto. - simpl. destruct a. destruct (f1 a). - simpl. simpl in IHp. destruct (transf_partial_program f1 p). - simpl. simpl in IHp. destruct (f2 b). - destruct (transf_partial_program f2 l). - rewrite <- IHp. auto. - rewrite <- IHp. auto. - auto. - simpl. rewrite <- IHp. simpl. destruct (f2 b); auto. - simpl. auto. -Qed. - -Lemma transform_partial_program_compose: - forall (A B C: Set) - (f1: A -> option B) (f2: B -> option C) - (p: program A), - transform_partial_program f1 p @@@ - (fun p' => transform_partial_program f2 p') = - transform_partial_program (fun f => f1 f @@@ f2) p. -Proof. - unfold transform_partial_program; intros. - rewrite <- transf_partial_program_compose. simpl. - destruct (transf_partial_program f1 (prog_funct p)); simpl. - auto. auto. -Qed. - -Lemma transf_program_partial_total: - forall (A B: Set) (f: A -> B) (l: list (ident * A)), - Some (AST.transf_program f l) = - AST.transf_partial_program (fun x => Some (f x)) l. -Proof. - induction l. simpl. auto. - simpl. destruct a. rewrite <- IHl. auto. -Qed. - -Lemma transform_program_partial_total: - forall (A B: Set) (f: A -> B) (p: program A), - Some (transform_program f p) = - transform_partial_program (fun x => Some (f x)) p. -Proof. - intros. unfold transform_program, transform_partial_program. - rewrite <- transf_program_partial_total. auto. -Qed. - -Lemma apply_total_transf_program: - forall (A B: Set) (f: A -> B) (x: option (program A)), - x @@ (fun p => transform_program f p) = - x @@@ (fun p => transform_partial_program (fun x => Some (f x)) p). -Proof. - intros. unfold apply_total, apply_partial. - destruct x. apply transform_program_partial_total. auto. -Qed. - -Lemma transf_cminor_program_equiv: - forall p, transf_cminor_program2 p = transf_cminor_program p. -Proof. - intro. unfold transf_cminor_program, transf_cminor_program2, transf_cminor_fundef. - simpl. - unfold RTLgen.transl_program, - Constprop.transf_program, RTL.program. - rewrite apply_total_transf_program. - rewrite transform_partial_program_compose. - unfold CSE.transf_program, RTL.program. - rewrite apply_total_transf_program. - rewrite transform_partial_program_compose. - unfold Allocation.transf_program, - LTL.program, RTL.program. - rewrite transform_partial_program_compose. - unfold Tunneling.tunnel_program, LTL.program. - rewrite apply_total_transf_program. - rewrite transform_partial_program_compose. - unfold Linearize.transf_program, LTL.program, Linear.program. - rewrite apply_total_transf_program. - rewrite transform_partial_program_compose. - unfold Stacking.transf_program, Linear.program, Mach.program. - rewrite transform_partial_program_compose. - unfold PPCgen.transf_program, Mach.program, PPC.program. - rewrite transform_partial_program_compose. - reflexivity. -Qed. - -Lemma transf_csharpminor_program_equiv: - forall p, transf_csharpminor_program2 p = transf_csharpminor_program p. -Proof. - intros. - unfold transf_csharpminor_program2, transf_csharpminor_program, transf_csharpminor_fundef. - simpl. - replace transf_cminor_program2 with transf_cminor_program. - unfold transf_cminor_program, Cminorgen.transl_program, Cminor.program, PPC.program. - apply transform_partial_program_compose. - symmetry. apply extensionality. exact transf_cminor_program_equiv. -Qed. - -(** * Semantic preservation *) - -(** We prove that the [transf_program2] translations preserve semantics. The proof - composes the semantic preservation results for each pass. *) - -Lemma transf_cminor_program2_correct: - forall p tp t n, - transf_cminor_program2 p = Some tp -> - Cminor.exec_program p t (Vint n) -> - PPC.exec_program tp t (Vint n). -Proof. - intros until n. - unfold transf_cminor_program2. - simpl. caseEq (RTLgen.transl_program p). intros p1 EQ1. - simpl. set (p2 := CSE.transf_program (Constprop.transf_program p1)). - caseEq (Allocation.transf_program p2). intros p3 EQ3. - simpl. set (p4 := Tunneling.tunnel_program p3). - set (p5 := Linearize.transf_program p4). - caseEq (Stacking.transf_program p5). intros p6 EQ6. - simpl. intros EQTP EXEC. - assert (WT3 : LTLtyping.wt_program p3). - apply Alloctyping.program_typing_preserved with p2. auto. - assert (WT4 : LTLtyping.wt_program p4). - unfold p4. apply Tunnelingtyping.program_typing_preserved. auto. - assert (WT5 : Lineartyping.wt_program p5). - unfold p5. apply Linearizetyping.program_typing_preserved. auto. - assert (WT6: Machtyping.wt_program p6). - apply Stackingtyping.program_typing_preserved with p5. auto. auto. - apply PPCgenproof.transf_program_correct with p6; auto. - apply Machabstr2mach.exec_program_equiv; auto. - apply Stackingproof.transl_program_correct with p5; auto. - unfold p5; apply Linearizeproof.transf_program_correct. - unfold p4; apply Tunnelingproof.transf_program_correct. - apply Allocproof.transl_program_correct with p2; auto. - unfold p2; apply CSEproof.transf_program_correct; - apply Constpropproof.transf_program_correct. - apply RTLgenproof.transl_program_correct with p; auto. - simpl; intros; discriminate. - simpl; intros; discriminate. - simpl; intros; discriminate. -Qed. - -Lemma transf_csharpminor_program2_correct: - forall p tp t n, - transf_csharpminor_program2 p = Some tp -> - Csharpminor.exec_program p t (Vint n) -> - PPC.exec_program tp t (Vint n). -Proof. - intros until n; unfold transf_csharpminor_program2; simpl. - caseEq (Cminorgen.transl_program p). - simpl; intros p1 EQ1 EQ2 EXEC. - apply transf_cminor_program2_correct with p1. auto. - apply Cminorgenproof.transl_program_correct with p. auto. - assumption. - simpl; intros; discriminate. -Qed. - -(** It follows that the whole compiler is semantic-preserving. *) - -Theorem transf_cminor_program_correct: - forall p tp t n, - transf_cminor_program p = Some tp -> - Cminor.exec_program p t (Vint n) -> - PPC.exec_program tp t (Vint n). -Proof. - intros. apply transf_cminor_program2_correct with p. - rewrite transf_cminor_program_equiv. assumption. assumption. -Qed. - -Theorem transf_csharpminor_program_correct: - forall p tp t n, - transf_csharpminor_program p = Some tp -> - Csharpminor.exec_program p t (Vint n) -> - PPC.exec_program tp t (Vint n). -Proof. - intros. apply transf_csharpminor_program2_correct with p. - rewrite transf_csharpminor_program_equiv. assumption. assumption. -Qed. diff --git a/backend/Mem.v b/backend/Mem.v deleted file mode 100644 index 7af696e1..00000000 --- a/backend/Mem.v +++ /dev/null @@ -1,2385 +0,0 @@ -(** This file develops the memory model that is used in the dynamic - semantics of all the languages of the compiler back-end. - It defines a type [mem] of memory states, the following 4 basic - operations over memory states, and their properties: -- [alloc]: allocate a fresh memory block; -- [free]: invalidate a memory block; -- [load]: read a memory chunk at a given address; -- [store]: store a memory chunk at a given address. -*) - -Require Import Coqlib. -Require Import Maps. -Require Import AST. -Require Import Integers. -Require Import Floats. -Require Import Values. - -(** * Structure of memory states *) - -(** A memory state is organized in several disjoint blocks. Each block - has a low and a high bound that defines its size. Each block map - byte offsets to the contents of this byte. *) - -Definition update (A: Set) (x: Z) (v: A) (f: Z -> A) : Z -> A := - fun y => if zeq y x then v else f y. - -Implicit Arguments update [A]. - -Lemma update_s: - forall (A: Set) (x: Z) (v: A) (f: Z -> A), - update x v f x = v. -Proof. - intros; unfold update. apply zeq_true. -Qed. - -Lemma update_o: - forall (A: Set) (x: Z) (v: A) (f: Z -> A) (y: Z), - x <> y -> update x v f y = f y. -Proof. - intros; unfold update. apply zeq_false; auto. -Qed. - -(** The possible contents of a byte-sized memory cell. To give intuitions, - a 4-byte value [v] stored at offset [d] will be represented by - the content [Datum32 v] at offset [d] and [Cont] at offsets [d+1], - [d+2] and [d+3]. The [Cont] contents enable detecting future writes - that would overlap partially the 4-byte value. *) - -Inductive content : Set := - | Undef: content (**r undefined contents *) - | Datum8: val -> content (**r a value that fits in 1 byte *) - | Datum16: val -> content (**r first byte of a 2-byte value *) - | Datum32: val -> content (**r first byte of a 4-byte value *) - | Datum64: val -> content (**r first byte of a 8-byte value *) - | Cont: content. (**r continuation bytes for a multi-byte value *) - -Definition contentmap : Set := Z -> content. - -(** A memory block comprises the dimensions of the block (low and high bounds) - plus a mapping from byte offsets to contents. For technical reasons, - we also carry around a proof that the mapping is equal to [Undef] - outside the range of allowed byte offsets. *) - -Record block_contents : Set := mkblock { - low: Z; - high: Z; - contents: contentmap; - undef_outside: forall ofs, ofs < low \/ ofs >= high -> contents ofs = Undef -}. - -(** A memory state is a mapping from block addresses (represented by [Z] - integers) to blocks. We also maintain the address of the next - unallocated block, and a proof that this address is positive. *) - -Record mem : Set := mkmem { - blocks: Z -> block_contents; - nextblock: block; - nextblock_pos: nextblock > 0 -}. - -(** * Operations on memory stores *) - -(** Memory reads and writes are performed by quantities called memory chunks, - encoding the type, size and signedness of the chunk being addressed. - It is useful to extract only the size information as given by the - following [memory_size] type. *) - -Inductive memory_size : Set := - | Size8: memory_size - | Size16: memory_size - | Size32: memory_size - | Size64: memory_size. - -Definition size_mem (sz: memory_size) := - match sz with - | Size8 => 1 - | Size16 => 2 - | Size32 => 4 - | Size64 => 8 - end. - -Definition mem_chunk (chunk: memory_chunk) := - match chunk with - | Mint8signed => Size8 - | Mint8unsigned => Size8 - | Mint16signed => Size16 - | Mint16unsigned => Size16 - | Mint32 => Size32 - | Mfloat32 => Size32 - | Mfloat64 => Size64 - end. - -Definition size_chunk (chunk: memory_chunk) := size_mem (mem_chunk chunk). - -(** The initial store. *) - -Remark one_pos: 1 > 0. -Proof. omega. Qed. - -Remark undef_undef_outside: - forall lo hi ofs, ofs < lo \/ ofs >= hi -> (fun y => Undef) ofs = Undef. -Proof. - auto. -Qed. - -Definition empty_block (lo hi: Z) : block_contents := - mkblock lo hi (fun y => Undef) (undef_undef_outside lo hi). - -Definition empty: mem := - mkmem (fun x => empty_block 0 0) 1 one_pos. - -Definition nullptr: block := 0. - -(** Allocation of a fresh block with the given bounds. Return an updated - memory state and the address of the fresh block, which initially contains - undefined cells. Note that allocation never fails: we model an - infinite memory. *) - -Remark succ_nextblock_pos: - forall m, Zsucc m.(nextblock) > 0. -Proof. intro. generalize (nextblock_pos m). omega. Qed. - -Definition alloc (m: mem) (lo hi: Z) := - (mkmem (update m.(nextblock) - (empty_block lo hi) - m.(blocks)) - (Zsucc m.(nextblock)) - (succ_nextblock_pos m), - m.(nextblock)). - -(** Freeing a block. Return the updated memory state where the given - block address has been invalidated: future reads and writes to this - address will fail. Note that we make no attempt to return the block - to an allocation pool: the given block address will never be allocated - later. *) - -Definition free (m: mem) (b: block) := - mkmem (update b - (empty_block 0 0) - m.(blocks)) - m.(nextblock) - m.(nextblock_pos). - -(** Freeing of a list of blocks. *) - -Definition free_list (m:mem) (l:list block) := - List.fold_right (fun b m => free m b) m l. - -(** Return the low and high bounds for the given block address. - Those bounds are 0 for freed or not yet allocated address. *) - -Definition low_bound (m: mem) (b: block) := - low (m.(blocks) b). -Definition high_bound (m: mem) (b: block) := - high (m.(blocks) b). - -(** A block address is valid if it was previously allocated. It remains valid - even after being freed. *) - -Definition valid_block (m: mem) (b: block) := - b < m.(nextblock). - -(** Reading and writing [N] adjacent locations in a [contentmap]. - - We define two functions and prove some of their properties: - - [getN n ofs m] returns the datum at offset [ofs] in block contents [m] - after checking that the contents of offsets [ofs+1] to [ofs+n+1] - are [Cont]. - - [setN n ofs v m] updates the block contents [m], storing the content [v] - at offset [ofs] and the content [Cont] at offsets [ofs+1] to [ofs+n+1]. - *) - -Fixpoint check_cont (n: nat) (p: Z) (m: contentmap) {struct n} : bool := - match n with - | O => true - | S n1 => - match m p with - | Cont => check_cont n1 (p + 1) m - | _ => false - end - end. - -Definition getN (n: nat) (p: Z) (m: contentmap) : content := - if check_cont n (p + 1) m then m p else Undef. - -Fixpoint set_cont (n: nat) (p: Z) (m: contentmap) {struct n} : contentmap := - match n with - | O => m - | S n1 => update p Cont (set_cont n1 (p + 1) m) - end. - -Definition setN (n: nat) (p: Z) (v: content) (m: contentmap) : contentmap := - update p v (set_cont n (p + 1) m). - -Lemma check_cont_true: - forall n p m, - (forall q, p <= q < p + Z_of_nat n -> m q = Cont) -> - check_cont n p m = true. -Proof. - induction n; intros. - reflexivity. - simpl. rewrite H. apply IHn. - intros. apply H. rewrite inj_S. omega. - rewrite inj_S. omega. -Qed. - -Hint Resolve check_cont_true. - -Lemma check_cont_inv: - forall n p m, - check_cont n p m = true -> - (forall q, p <= q < p + Z_of_nat n -> m q = Cont). -Proof. - induction n; intros until m. - unfold Z_of_nat. intros. omegaContradiction. - unfold check_cont; fold check_cont. - caseEq (m p); intros; try discriminate. - assert (p = q \/ p + 1 <= q < (p + 1) + Z_of_nat n). - rewrite inj_S in H1. omega. - elim H2; intro. - subst q. auto. - apply IHn with (p + 1); auto. -Qed. - -Hint Resolve check_cont_inv. - -Lemma check_cont_false: - forall n p q m, - p <= q < p + Z_of_nat n -> - m q <> Cont -> - check_cont n p m = false. -Proof. - intros. caseEq (check_cont n p m); intro. - generalize (check_cont_inv _ _ _ H1 q H). intro. contradiction. - auto. -Qed. - -Hint Resolve check_cont_false. - -Lemma set_cont_inside: - forall n p m q, - p <= q < p + Z_of_nat n -> - (set_cont n p m) q = Cont. -Proof. - induction n; intros. - unfold Z_of_nat in H. omegaContradiction. - simpl. - assert (p = q \/ p + 1 <= q < (p + 1) + Z_of_nat n). - rewrite inj_S in H. omega. - elim H0; intro. - subst q. apply update_s. - rewrite update_o. apply IHn. auto. - red; intro; subst q. omega. -Qed. - -Hint Resolve set_cont_inside. - -Lemma set_cont_outside: - forall n p m q, - q < p \/ p + Z_of_nat n <= q -> - (set_cont n p m) q = m q. -Proof. - induction n; intros. - simpl. auto. - simpl. rewrite inj_S in H. - rewrite update_o. apply IHn. omega. omega. -Qed. - -Hint Resolve set_cont_outside. - -Lemma getN_setN_same: - forall n p v m, - getN n p (setN n p v m) = v. -Proof. - intros. unfold getN, setN. - rewrite check_cont_true. apply update_s. - intros. rewrite update_o. apply set_cont_inside. auto. - omega. -Qed. - -Hint Resolve getN_setN_same. - -Lemma getN_setN_other: - forall n1 n2 p1 p2 v m, - p1 + Z_of_nat n1 < p2 \/ p2 + Z_of_nat n2 < p1 -> - getN n2 p2 (setN n1 p1 v m) = getN n2 p2 m. -Proof. - intros. unfold getN, setN. - caseEq (check_cont n2 (p2 + 1) m); intro. - rewrite check_cont_true. rewrite update_o. - apply set_cont_outside. omega. omega. - intros. rewrite update_o. rewrite set_cont_outside. - eapply check_cont_inv. eauto. auto. - omega. omega. - caseEq (check_cont n2 (p2 + 1) (update p1 v (set_cont n1 (p1 + 1) m))); intros. - assert (check_cont n2 (p2 + 1) m = true). - apply check_cont_true. intros. - generalize (check_cont_inv _ _ _ H1 q H2). - rewrite update_o. rewrite set_cont_outside. auto. omega. omega. - rewrite H0 in H2; discriminate. - auto. -Qed. - -Hint Resolve getN_setN_other. - -Lemma getN_setN_overlap: - forall n1 n2 p1 p2 v m, - p1 <> p2 -> - p1 + Z_of_nat n1 >= p2 -> p2 + Z_of_nat n2 >= p1 -> - v <> Cont -> - getN n2 p2 (setN n1 p1 v m) = Cont \/ - getN n2 p2 (setN n1 p1 v m) = Undef. -Proof. - intros. unfold getN. - caseEq (check_cont n2 (p2 + 1) (setN n1 p1 v m)); intro. - case (zlt p2 p1); intro. - assert (p2 + 1 <= p1 < p2 + 1 + Z_of_nat n2). omega. - generalize (check_cont_inv _ _ _ H3 p1 H4). - unfold setN. rewrite update_s. intro. contradiction. - left. unfold setN. rewrite update_o. - apply set_cont_inside. omega. auto. - right; auto. -Qed. - -Hint Resolve getN_setN_overlap. - -Lemma getN_setN_mismatch: - forall n1 n2 p v m, - getN n2 p (setN n1 p v m) = v \/ getN n2 p (setN n1 p v m) = Undef. -Proof. - intros. unfold getN. - caseEq (check_cont n2 (p + 1) (setN n1 p v m)); intro. - left. unfold setN. apply update_s. - right. auto. -Qed. - -Hint Resolve getN_setN_mismatch. - -Lemma getN_init: - forall n p, - getN n p (fun y => Undef) = Undef. -Proof. - intros. unfold getN. - case (check_cont n (p + 1) (fun y => Undef)); auto. -Qed. - -Hint Resolve getN_init. - -(** The following function checks whether accessing the given memory chunk - at the given offset in the given block respects the bounds of the block. *) - -Definition in_bounds (chunk: memory_chunk) (ofs: Z) (c: block_contents) : - {c.(low) <= ofs /\ ofs + size_chunk chunk <= c.(high)} - + {c.(low) > ofs \/ ofs + size_chunk chunk > c.(high)} := - match zle c.(low) ofs, zle (ofs + size_chunk chunk) c.(high) with - | left P1, left P2 => left _ (conj P1 P2) - | left P1, right P2 => right _ (or_intror _ P2) - | right P1, _ => right _ (or_introl _ P1) - end. - -Lemma in_bounds_holds: - forall (chunk: memory_chunk) (ofs: Z) (c: block_contents) - (A: Set) (a b: A), - c.(low) <= ofs -> ofs + size_chunk chunk <= c.(high) -> - (if in_bounds chunk ofs c then a else b) = a. -Proof. - intros. case (in_bounds chunk ofs c); intro. - auto. - omegaContradiction. -Qed. - -Lemma in_bounds_exten: - forall (chunk: memory_chunk) (ofs: Z) (c: block_contents) (x: contentmap) P, - in_bounds chunk ofs (mkblock (low c) (high c) x P) = - in_bounds chunk ofs c. -Proof. - intros; reflexivity. -Qed. - -Hint Resolve in_bounds_holds in_bounds_exten. - -(** [valid_pointer] holds if the given block address is valid and the - given offset falls within the bounds of the corresponding block. *) - -Definition valid_pointer (m: mem) (b: block) (ofs: Z) : bool := - if zlt b m.(nextblock) then - (let c := m.(blocks) b in - if zle c.(low) ofs then if zlt ofs c.(high) then true else false - else false) - else false. - -(** Read a quantity of size [sz] at offset [ofs] in block contents [c]. - Return [Vundef] if the requested size does not match that of the - current contents, or if the following offsets do not contain [Cont]. - The first check captures a size mismatch between the read and the - latest write at this offset. The second check captures partial overwriting - of the latest write at this offset by a more recent write at a nearby - offset. *) - -Definition load_contents (sz: memory_size) (c: contentmap) (ofs: Z) : val := - match sz with - | Size8 => - match getN 0%nat ofs c with - | Datum8 n => n - | _ => Vundef - end - | Size16 => - match getN 1%nat ofs c with - | Datum16 n => n - | _ => Vundef - end - | Size32 => - match getN 3%nat ofs c with - | Datum32 n => n - | _ => Vundef - end - | Size64 => - match getN 7%nat ofs c with - | Datum64 n => n - | _ => Vundef - end - end. - -(** [load chunk m b ofs] perform a read in memory state [m], at address - [b] and offset [ofs]. [None] is returned if the address is invalid - or the memory access is out of bounds. *) - -Definition load (chunk: memory_chunk) (m: mem) (b: block) (ofs: Z) - : option val := - if zlt b m.(nextblock) then - (let c := m.(blocks) b in - if in_bounds chunk ofs c - then Some (Val.load_result chunk - (load_contents (mem_chunk chunk) c.(contents) ofs)) - else None) - else - None. - -(** [loadv chunk m addr] is similar, but the address and offset are given - as a single value [addr], which must be a pointer value. *) - -Definition loadv (chunk: memory_chunk) (m: mem) (addr: val) : option val := - match addr with - | Vptr b ofs => load chunk m b (Int.signed ofs) - | _ => None - end. - -Theorem load_in_bounds: - forall (chunk: memory_chunk) (m: mem) (b: block) (ofs: Z), - valid_block m b -> - low_bound m b <= ofs -> - ofs + size_chunk chunk <= high_bound m b -> - exists v, load chunk m b ofs = Some v. -Proof. - intros. unfold load. rewrite zlt_true; auto. - rewrite in_bounds_holds. - exists (Val.load_result chunk - (load_contents (mem_chunk chunk) - (contents (m.(blocks) b)) - ofs)). - auto. - exact H0. exact H1. -Qed. - -Lemma load_inv: - forall (chunk: memory_chunk) (m: mem) (b: block) (ofs: Z) (v: val), - load chunk m b ofs = Some v -> - let c := m.(blocks) b in - b < m.(nextblock) /\ - c.(low) <= ofs /\ - ofs + size_chunk chunk <= c.(high) /\ - Val.load_result chunk (load_contents (mem_chunk chunk) c.(contents) ofs) = v. -Proof. - intros until v; unfold load. - case (zlt b (nextblock m)); intro. - set (c := m.(blocks) b). - case (in_bounds chunk ofs c). - intuition congruence. - intros; discriminate. - intros; discriminate. -Qed. -Hint Resolve load_in_bounds load_inv. - -(* Write the value [v] with size [sz] at offset [ofs] in block contents [c]. - Return updated block contents. [Cont] contents are stored at the following - offsets. *) - -Definition store_contents (sz: memory_size) (c: contentmap) - (ofs: Z) (v: val) : contentmap := - match sz with - | Size8 => - setN 0%nat ofs (Datum8 v) c - | Size16 => - setN 1%nat ofs (Datum16 v) c - | Size32 => - setN 3%nat ofs (Datum32 v) c - | Size64 => - setN 7%nat ofs (Datum64 v) c - end. - -Remark store_contents_undef_outside: - forall sz c ofs v lo hi, - lo <= ofs /\ ofs + size_mem sz <= hi -> - (forall x, x < lo \/ x >= hi -> c x = Undef) -> - (forall x, x < lo \/ x >= hi -> - store_contents sz c ofs v x = Undef). -Proof. - intros until hi; intros [LO HI] UO. - assert (forall n d x, - ofs + (1 + Z_of_nat n) <= hi -> - x < lo \/ x >= hi -> - setN n ofs d c x = Undef). - intros. unfold setN. rewrite update_o. - transitivity (c x). apply set_cont_outside. omega. - apply UO. omega. omega. - unfold store_contents; destruct sz; unfold size_mem in HI; - intros; apply H; auto; simpl Z_of_nat; auto. -Qed. - -Definition unchecked_store - (chunk: memory_chunk) (m: mem) (b: block) - (ofs: Z) (v: val) - (P: (m.(blocks) b).(low) <= ofs /\ - ofs + size_chunk chunk <= (m.(blocks) b).(high)) : mem := - let c := m.(blocks) b in - mkmem - (update b - (mkblock c.(low) c.(high) - (store_contents (mem_chunk chunk) c.(contents) ofs v) - (store_contents_undef_outside (mem_chunk chunk) c.(contents) - ofs v _ _ P c.(undef_outside))) - m.(blocks)) - m.(nextblock) - m.(nextblock_pos). - -(** [store chunk m b ofs v] perform a write in memory state [m]. - Value [v] is stored at address [b] and offset [ofs]. - Return the updated memory store, or [None] if the address is invalid - or the memory access is out of bounds. *) - -Definition store (chunk: memory_chunk) (m: mem) (b: block) - (ofs: Z) (v: val) : option mem := - if zlt b m.(nextblock) then - match in_bounds chunk ofs (m.(blocks) b) with - | left P => Some(unchecked_store chunk m b ofs v P) - | right _ => None - end - else - None. - -(** [storev chunk m addr v] is similar, but the address and offset are given - as a single value [addr], which must be a pointer value. *) - -Definition storev (chunk: memory_chunk) (m: mem) (addr v: val) : option mem := - match addr with - | Vptr b ofs => store chunk m b (Int.signed ofs) v - | _ => None - end. - -Theorem store_in_bounds: - forall (chunk: memory_chunk) (m: mem) (b: block) (ofs: Z) (v: val), - valid_block m b -> - low_bound m b <= ofs -> - ofs + size_chunk chunk <= high_bound m b -> - exists m', store chunk m b ofs v = Some m'. -Proof. - intros. unfold store. - rewrite zlt_true; auto. - case (in_bounds chunk ofs (blocks m b)); intro P. - exists (unchecked_store chunk m b ofs v P). reflexivity. - unfold low_bound in H0. unfold high_bound in H1. omegaContradiction. -Qed. - -Lemma store_inv: - forall (chunk: memory_chunk) (m m': mem) (b: block) (ofs: Z) (v: val), - store chunk m b ofs v = Some m' -> - let c := m.(blocks) b in - b < m.(nextblock) /\ - c.(low) <= ofs /\ - ofs + size_chunk chunk <= c.(high) /\ - m'.(nextblock) = m.(nextblock) /\ - exists P, m'.(blocks) = - update b (mkblock c.(low) c.(high) - (store_contents (mem_chunk chunk) c.(contents) ofs v) P) - m.(blocks). -Proof. - intros until v; unfold store. - case (zlt b (nextblock m)); intro. - set (c := m.(blocks) b). - case (in_bounds chunk ofs c). - intros; injection H; intro; subst m'. simpl. - intuition. fold c. - exists (store_contents_undef_outside (mem_chunk chunk) - (contents c) ofs v (low c) (high c) a (undef_outside c)). - auto. - intros; discriminate. - intros; discriminate. -Qed. - -Hint Resolve store_in_bounds store_inv. - -(** Build a block filled with the given initialization data. *) - -Fixpoint contents_init_data (pos: Z) (id: list init_data) {struct id}: contentmap := - match id with - | nil => (fun y => Undef) - | Init_int8 n :: id' => - store_contents Size8 (contents_init_data (pos + 1) id') pos (Vint n) - | Init_int16 n :: id' => - store_contents Size16 (contents_init_data (pos + 2) id') pos (Vint n) - | Init_int32 n :: id' => - store_contents Size32 (contents_init_data (pos + 4) id') pos (Vint n) - | Init_float32 f :: id' => - store_contents Size32 (contents_init_data (pos + 4) id') pos (Vfloat f) - | Init_float64 f :: id' => - store_contents Size64 (contents_init_data (pos + 8) id') pos (Vfloat f) - | Init_space n :: id' => - contents_init_data (pos + Zmax n 0) id' - end. - -Definition size_init_data (id: init_data) : Z := - match id with - | Init_int8 _ => 1 - | Init_int16 _ => 2 - | Init_int32 _ => 4 - | Init_float32 _ => 4 - | Init_float64 _ => 8 - | Init_space n => Zmax n 0 - end. - -Definition size_init_data_list (id: list init_data): Z := - List.fold_right (fun id sz => size_init_data id + sz) 0 id. - -Remark size_init_data_list_pos: - forall id, size_init_data_list id >= 0. -Proof. - induction id; simpl. - omega. - assert (size_init_data a >= 0). destruct a; simpl; try omega. - generalize (Zmax2 z 0). omega. omega. -Qed. - -Remark contents_init_data_undef_outside: - forall id pos x, - x < pos \/ x >= pos + size_init_data_list id -> - contents_init_data pos id x = Undef. -Proof. - induction id; simpl; intros. - auto. - generalize (size_init_data_list_pos id); intro. - assert (forall n delta dt, - delta = 1 + Z_of_nat n -> - x < pos \/ x >= pos + (delta + size_init_data_list id) -> - setN n pos dt (contents_init_data (pos + delta) id) x = Undef). - intros. unfold setN. rewrite update_o. - transitivity (contents_init_data (pos + delta) id x). - apply set_cont_outside. omega. - apply IHid. omega. omega. - unfold size_init_data in H; destruct a; - try (apply H1; [reflexivity|assumption]). - apply IHid. generalize (Zmax2 z 0). omega. -Qed. - -Definition block_init_data (id: list init_data) : block_contents := - mkblock 0 (size_init_data_list id) - (contents_init_data 0 id) - (contents_init_data_undef_outside id 0). - -Definition alloc_init_data (m: mem) (id: list init_data) : mem * block := - (mkmem (update m.(nextblock) - (block_init_data id) - m.(blocks)) - (Zsucc m.(nextblock)) - (succ_nextblock_pos m), - m.(nextblock)). - -Remark block_init_data_empty: - block_init_data nil = empty_block 0 0. -Proof. - unfold block_init_data, empty_block. simpl. - decEq. apply proof_irrelevance. -Qed. - -(** * Properties of the memory operations *) - -(** ** Properties related to block validity *) - -Lemma valid_not_valid_diff: - forall m b b', valid_block m b -> ~(valid_block m b') -> b <> b'. -Proof. - intros; red; intros. subst b'. contradiction. -Qed. - -Theorem fresh_block_alloc: - forall (m1 m2: mem) (lo hi: Z) (b: block), - alloc m1 lo hi = (m2, b) -> ~(valid_block m1 b). -Proof. - intros. injection H; intros; subst b. - unfold valid_block. omega. -Qed. - -Theorem valid_new_block: - forall (m1 m2: mem) (lo hi: Z) (b: block), - alloc m1 lo hi = (m2, b) -> valid_block m2 b. -Proof. - unfold alloc, valid_block; intros. - injection H; intros. subst b; subst m2; simpl. omega. -Qed. - -Theorem valid_block_alloc: - forall (m1 m2: mem) (lo hi: Z) (b b': block), - alloc m1 lo hi = (m2, b') -> - valid_block m1 b -> valid_block m2 b. -Proof. - unfold alloc, valid_block; intros. - injection H; intros. subst m2; simpl. omega. -Qed. - -Theorem valid_block_store: - forall (chunk: memory_chunk) (m1 m2: mem) (b b': block) (ofs: Z) (v: val), - store chunk m1 b' ofs v = Some m2 -> - valid_block m1 b -> valid_block m2 b. -Proof. - intros. generalize (store_inv _ _ _ _ _ _ H). - intros [A [B [C [D [P E]]]]]. - red. rewrite D. exact H0. -Qed. - -Theorem valid_block_free: - forall (m: mem) (b b': block), - valid_block m b -> valid_block (free m b') b. -Proof. - unfold valid_block, free; intros. - simpl. auto. -Qed. - -(** ** Properties related to [alloc] *) - -Theorem load_alloc_other: - forall (chunk: memory_chunk) (m1 m2: mem) - (b b': block) (ofs lo hi: Z) (v: val), - alloc m1 lo hi = (m2, b') -> - load chunk m1 b ofs = Some v -> - load chunk m2 b ofs = Some v. -Proof. - unfold alloc; intros. - injection H; intros; subst m2; clear H. - generalize (load_inv _ _ _ _ _ H0). - intros (A, (B, (C, D))). - unfold load; simpl. - rewrite zlt_true. - repeat (rewrite update_o). - rewrite in_bounds_holds. congruence. auto. auto. - omega. omega. -Qed. - -Lemma load_contents_init: - forall (sz: memory_size) (ofs: Z), - load_contents sz (fun y => Undef) ofs = Vundef. -Proof. - intros. destruct sz; reflexivity. -Qed. - -Theorem load_alloc_same: - forall (chunk: memory_chunk) (m1 m2: mem) - (b b': block) (ofs lo hi: Z) (v: val), - alloc m1 lo hi = (m2, b') -> - load chunk m2 b' ofs = Some v -> - v = Vundef. -Proof. - unfold alloc; intros. - injection H; intros; subst m2; clear H. - generalize (load_inv _ _ _ _ _ H0). - simpl. rewrite H1. rewrite update_s. simpl. intros (A, (B, (C, D))). - rewrite <- D. rewrite load_contents_init. - destruct chunk; reflexivity. -Qed. - -Theorem low_bound_alloc: - forall (m1 m2: mem) (b b': block) (lo hi: Z), - alloc m1 lo hi = (m2, b') -> - low_bound m2 b = if zeq b b' then lo else low_bound m1 b. -Proof. - unfold alloc; intros. - injection H; intros; subst m2; clear H. - unfold low_bound; simpl. - unfold update. - subst b'. - case (zeq b (nextblock m1)); reflexivity. -Qed. - -Theorem high_bound_alloc: - forall (m1 m2: mem) (b b': block) (lo hi: Z), - alloc m1 lo hi = (m2, b') -> - high_bound m2 b = if zeq b b' then hi else high_bound m1 b. -Proof. - unfold alloc; intros. - injection H; intros; subst m2; clear H. - unfold high_bound; simpl. - unfold update. - subst b'. - case (zeq b (nextblock m1)); reflexivity. -Qed. - -Theorem store_alloc: - forall (chunk: memory_chunk) (m1 m2: mem) (b: block) (ofs lo hi: Z) (v: val), - alloc m1 lo hi = (m2, b) -> - lo <= ofs -> ofs + size_chunk chunk <= hi -> - exists m2', store chunk m2 b ofs v = Some m2'. -Proof. - unfold alloc; intros. - injection H; intros. - assert (A: b < m2.(nextblock)). - subst m2; subst b; simpl; omega. - assert (B: low_bound m2 b <= ofs). - subst m2; subst b. unfold low_bound; simpl. rewrite update_s. - simpl. assumption. - assert (C: ofs + size_chunk chunk <= high_bound m2 b). - subst m2; subst b. unfold high_bound; simpl. rewrite update_s. - simpl. assumption. - exact (store_in_bounds chunk m2 b ofs v A B C). -Qed. - -Hint Resolve store_alloc high_bound_alloc low_bound_alloc load_alloc_same -load_contents_init load_alloc_other. - -(** ** Properties related to [free] *) - -Theorem load_free: - forall (chunk: memory_chunk) (m: mem) (b bf: block) (ofs: Z), - b <> bf -> - load chunk (free m bf) b ofs = load chunk m b ofs. -Proof. - intros. unfold free, load; simpl. - case (zlt b (nextblock m)). - repeat (rewrite update_o; auto). - reflexivity. -Qed. - -Theorem low_bound_free: - forall (m: mem) (b bf: block), - b <> bf -> - low_bound (free m bf) b = low_bound m b. -Proof. - intros. unfold free, low_bound; simpl. - rewrite update_o; auto. -Qed. - -Theorem high_bound_free: - forall (m: mem) (b bf: block), - b <> bf -> - high_bound (free m bf) b = high_bound m b. -Proof. - intros. unfold free, high_bound; simpl. - rewrite update_o; auto. -Qed. -Hint Resolve load_free low_bound_free high_bound_free. - -(** ** Properties related to [store] *) - -Lemma store_is_in_bounds: - forall chunk m1 b ofs v m2, - store chunk m1 b ofs v = Some m2 -> - low_bound m1 b <= ofs /\ ofs + size_chunk chunk <= high_bound m1 b. -Proof. - intros. generalize (store_inv _ _ _ _ _ _ H). - intros [A [B [C [P D]]]]. - unfold low_bound, high_bound. tauto. -Qed. - -Lemma load_store_contents_same: - forall (sz: memory_size) (c: contentmap) (ofs: Z) (v: val), - load_contents sz (store_contents sz c ofs v) ofs = v. -Proof. - intros until v. - unfold load_contents, store_contents in |- *; case sz; - rewrite getN_setN_same; reflexivity. -Qed. - -Theorem load_store_same: - forall (chunk: memory_chunk) (m1 m2: mem) (b: block) (ofs: Z) (v: val), - store chunk m1 b ofs v = Some m2 -> - load chunk m2 b ofs = Some (Val.load_result chunk v). -Proof. - intros. - generalize (store_inv _ _ _ _ _ _ H). - intros (A, (B, (C, (D, (P, E))))). - unfold load. rewrite D. - rewrite zlt_true; auto. rewrite E. - repeat (rewrite update_s). simpl. - rewrite in_bounds_exten. rewrite in_bounds_holds; auto. - rewrite load_store_contents_same; auto. -Qed. - -Lemma load_store_contents_other: - forall (sz1 sz2: memory_size) (c: contentmap) - (ofs1 ofs2: Z) (v: val), - ofs2 + size_mem sz2 <= ofs1 \/ ofs1 + size_mem sz1 <= ofs2 -> - load_contents sz2 (store_contents sz1 c ofs1 v) ofs2 = - load_contents sz2 c ofs2. -Proof. - intros until v. - unfold size_mem, store_contents, load_contents; - case sz1; case sz2; intros; - (rewrite getN_setN_other; - [reflexivity | simpl Z_of_nat; omega]). -Qed. - -Theorem load_store_other: - forall (chunk1 chunk2: memory_chunk) (m1 m2: mem) - (b1 b2: block) (ofs1 ofs2: Z) (v: val), - store chunk1 m1 b1 ofs1 v = Some m2 -> - b1 <> b2 - \/ ofs2 + size_chunk chunk2 <= ofs1 - \/ ofs1 + size_chunk chunk1 <= ofs2 -> - load chunk2 m2 b2 ofs2 = load chunk2 m1 b2 ofs2. -Proof. - intros. - generalize (store_inv _ _ _ _ _ _ H). - intros (A, (B, (C, (D, (P, E))))). - unfold load. rewrite D. - case (zlt b2 (nextblock m1)); intro. - rewrite E; unfold update; case (zeq b2 b1); intro; simpl. - subst b2. rewrite in_bounds_exten. - rewrite load_store_contents_other. auto. - tauto. - reflexivity. - reflexivity. -Qed. - -Ltac LSCO := - match goal with - | |- (match getN ?sz2 ?ofs2 (setN ?sz1 ?ofs1 ?v ?c) with - | Undef => _ - | Datum8 _ => _ - | Datum16 _ => _ - | Datum32 _ => _ - | Datum64 _ => _ - | Cont => _ - end = _) => - elim (getN_setN_overlap sz1 sz2 ofs1 ofs2 v c); - [ let H := fresh in (intro H; rewrite H; reflexivity) - | let H := fresh in (intro H; rewrite H; reflexivity) - | assumption - | simpl Z_of_nat; omega - | simpl Z_of_nat; omega - | discriminate ] - end. - -Lemma load_store_contents_overlap: - forall (sz1 sz2: memory_size) (c: contentmap) - (ofs1 ofs2: Z) (v: val), - ofs1 <> ofs2 -> - ofs2 + size_mem sz2 > ofs1 -> ofs1 + size_mem sz1 > ofs2 -> - load_contents sz2 (store_contents sz1 c ofs1 v) ofs2 = Vundef. -Proof. - intros. - destruct sz1; destruct sz2; simpl in H0; simpl in H1; simpl; LSCO. -Qed. - -Ltac LSCM := - match goal with - | H:(?x <> ?x) |- _ => - elim H; reflexivity - | |- (match getN ?sz2 ?ofs (setN ?sz1 ?ofs ?v ?c) with - | Undef => _ - | Datum8 _ => _ - | Datum16 _ => _ - | Datum32 _ => _ - | Datum64 _ => _ - | Cont => _ - end = _) => - elim (getN_setN_mismatch sz1 sz2 ofs v c); - [ let H := fresh in (intro H; rewrite H; reflexivity) - | let H := fresh in (intro H; rewrite H; reflexivity) ] - end. - -Lemma load_store_contents_mismatch: - forall (sz1 sz2: memory_size) (c: contentmap) - (ofs: Z) (v: val), - sz1 <> sz2 -> - load_contents sz2 (store_contents sz1 c ofs v) ofs = Vundef. -Proof. - intros. - destruct sz1; destruct sz2; simpl; LSCM. -Qed. - -Theorem low_bound_store: - forall (chunk: memory_chunk) (m1 m2: mem) (b b': block) (ofs: Z) (v: val), - store chunk m1 b ofs v = Some m2 -> - low_bound m2 b' = low_bound m1 b'. -Proof. - intros. - generalize (store_inv _ _ _ _ _ _ H). - intros (A, (B, (C, (D, (P, E))))). - unfold low_bound. rewrite E; unfold update. - case (zeq b' b); intro. - subst b'. reflexivity. - reflexivity. -Qed. - -Theorem high_bound_store: - forall (chunk: memory_chunk) (m1 m2: mem) (b b': block) (ofs: Z) (v: val), - store chunk m1 b ofs v = Some m2 -> - high_bound m2 b' = high_bound m1 b'. -Proof. - intros. - generalize (store_inv _ _ _ _ _ _ H). - intros (A, (B, (C, (D, (P, E))))). - unfold high_bound. rewrite E; unfold update. - case (zeq b' b); intro. - subst b'. reflexivity. - reflexivity. -Qed. - -Hint Resolve high_bound_store low_bound_store load_store_contents_mismatch - load_store_contents_overlap load_store_other store_is_in_bounds - load_store_contents_same load_store_same load_store_contents_other. - -(** * Agreement between memory blocks. *) - -(** Two memory blocks [c1] and [c2] agree on a range [lo] to [hi] - if they associate the same contents to byte offsets in the range - [lo] (included) to [hi] (excluded). *) - -Definition contentmap_agree (lo hi: Z) (c1 c2: contentmap) := - forall (ofs: Z), - lo <= ofs < hi -> c1 ofs = c2 ofs. - -Definition block_contents_agree (lo hi: Z) (c1 c2: block_contents) := - contentmap_agree lo hi c1.(contents) c2.(contents). - -Definition block_agree (b: block) (lo hi: Z) (m1 m2: mem) := - block_contents_agree lo hi - (m1.(blocks) b) (m2.(blocks) b). - -Theorem block_agree_refl: - forall (m: mem) (b: block) (lo hi: Z), - block_agree b lo hi m m. -Proof. - intros. hnf. auto. -Qed. - -Theorem block_agree_sym: - forall (m1 m2: mem) (b: block) (lo hi: Z), - block_agree b lo hi m1 m2 -> - block_agree b lo hi m2 m1. -Proof. - intros. hnf. intros. symmetry. apply H. auto. -Qed. - -Theorem block_agree_trans: - forall (m1 m2 m3: mem) (b: block) (lo hi: Z), - block_agree b lo hi m1 m2 -> - block_agree b lo hi m2 m3 -> - block_agree b lo hi m1 m3. -Proof. - intros. hnf. intros. - transitivity (contents (blocks m2 b) ofs). - apply H; auto. apply H0; auto. -Qed. - -Lemma check_cont_agree: - forall (c1 c2: contentmap) (lo hi: Z), - contentmap_agree lo hi c1 c2 -> - forall (n: nat) (ofs: Z), - lo <= ofs -> ofs + Z_of_nat n <= hi -> - check_cont n ofs c2 = check_cont n ofs c1. -Proof. - induction n; intros; simpl. - auto. - rewrite inj_S in H1. - rewrite H. case (c2 ofs); intros; auto. - apply IHn; omega. - omega. -Qed. - -Lemma getN_agree: - forall (c1 c2: contentmap) (lo hi: Z), - contentmap_agree lo hi c1 c2 -> - forall (n: nat) (ofs: Z), - lo <= ofs -> ofs + Z_of_nat n < hi -> - getN n ofs c2 = getN n ofs c1. -Proof. - intros. unfold getN. - rewrite (check_cont_agree c1 c2 lo hi H n (ofs + 1)). - case (check_cont n (ofs + 1) c1). - symmetry. apply H. omega. auto. - omega. omega. -Qed. - -Lemma load_contentmap_agree: - forall (sz: memory_size) (c1 c2: contentmap) (lo hi ofs: Z), - contentmap_agree lo hi c1 c2 -> - lo <= ofs -> ofs + size_mem sz <= hi -> - load_contents sz c2 ofs = load_contents sz c1 ofs. -Proof. - intros sz c1 c2 lo hi ofs EX LO. - unfold load_contents, size_mem; case sz; intro HI; - rewrite (getN_agree c1 c2 lo hi EX); auto; simpl Z_of_nat; omega. -Qed. - -Lemma set_cont_agree: - forall (lo hi: Z) (n: nat) (c1 c2: contentmap) (ofs: Z), - contentmap_agree lo hi c1 c2 -> - contentmap_agree lo hi (set_cont n ofs c1) (set_cont n ofs c2). -Proof. - induction n; simpl; intros. - auto. - red. intros p B. - case (zeq p ofs); intro. - subst p. repeat (rewrite update_s). reflexivity. - repeat (rewrite update_o). apply IHn. assumption. - assumption. auto. auto. -Qed. - -Lemma setN_agree: - forall (lo hi: Z) (n: nat) (c1 c2: contentmap) (ofs: Z) (v: content), - contentmap_agree lo hi c1 c2 -> - contentmap_agree lo hi (setN n ofs v c1) (setN n ofs v c2). -Proof. - intros. unfold setN. red; intros p B. - case (zeq p ofs); intro. - subst p. repeat (rewrite update_s). reflexivity. - repeat (rewrite update_o; auto). - exact (set_cont_agree lo hi n c1 c2 (ofs + 1) H p B). -Qed. - -Lemma store_contentmap_agree: - forall (sz: memory_size) (c1 c2: contentmap) (lo hi ofs: Z) (v: val), - contentmap_agree lo hi c1 c2 -> - contentmap_agree lo hi - (store_contents sz c1 ofs v) (store_contents sz c2 ofs v). -Proof. - intros. unfold store_contents; case sz; apply setN_agree; auto. -Qed. - -Lemma set_cont_outside_agree: - forall (lo hi: Z) (n: nat) (c1 c2: contentmap) (ofs: Z), - contentmap_agree lo hi c1 c2 -> - ofs + Z_of_nat n <= lo \/ hi <= ofs -> - contentmap_agree lo hi c1 (set_cont n ofs c2). -Proof. - induction n; intros; simpl. - auto. - red; intros p R. rewrite inj_S in H0. - unfold update. case (zeq p ofs); intro. - subst p. omegaContradiction. - apply IHn. auto. omega. auto. -Qed. - -Lemma setN_outside_agree: - forall (lo hi: Z) (n: nat) (c1 c2: contentmap) (ofs: Z) (v: content), - contentmap_agree lo hi c1 c2 -> - ofs + Z_of_nat n < lo \/ hi <= ofs -> - contentmap_agree lo hi c1 (setN n ofs v c2). -Proof. - intros. unfold setN. red; intros p R. - unfold update. case (zeq p ofs); intro. - omegaContradiction. - apply (set_cont_outside_agree lo hi n c1 c2 (ofs + 1) H). - omega. auto. -Qed. - -Lemma store_contentmap_outside_agree: - forall (sz: memory_size) (c1 c2: contentmap) (lo hi ofs: Z) (v: val), - contentmap_agree lo hi c1 c2 -> - ofs + size_mem sz <= lo \/ hi <= ofs -> - contentmap_agree lo hi c1 (store_contents sz c2 ofs v). -Proof. - intros until v. - unfold store_contents; case sz; simpl; intros; - apply setN_outside_agree; auto; simpl Z_of_nat; omega. -Qed. - -Theorem load_agree: - forall (chunk: memory_chunk) (m1 m2: mem) - (b: block) (lo hi: Z) (ofs: Z) (v1 v2: val), - block_agree b lo hi m1 m2 -> - lo <= ofs -> ofs + size_chunk chunk <= hi -> - load chunk m1 b ofs = Some v1 -> - load chunk m2 b ofs = Some v2 -> - v1 = v2. -Proof. - intros. - generalize (load_inv _ _ _ _ _ H2). intros [K [L [M N]]]. - generalize (load_inv _ _ _ _ _ H3). intros [P [Q [R S]]]. - subst v1; subst v2. symmetry. - decEq. eapply load_contentmap_agree. - apply H. auto. auto. -Qed. - -Theorem store_agree: - forall (chunk: memory_chunk) (m1 m2 m1' m2': mem) - (b: block) (lo hi: Z) - (b': block) (ofs: Z) (v: val), - block_agree b lo hi m1 m2 -> - store chunk m1 b' ofs v = Some m1' -> - store chunk m2 b' ofs v = Some m2' -> - block_agree b lo hi m1' m2'. -Proof. - intros. - generalize (store_inv _ _ _ _ _ _ H0). - intros [I [J [K [L [x M]]]]]. - generalize (store_inv _ _ _ _ _ _ H1). - intros [P [Q [R [S [y T]]]]]. - red. red. - rewrite M; rewrite T; unfold update. - case (zeq b b'); intro; simpl. - subst b'. apply store_contentmap_agree. assumption. - apply H. -Qed. - -Theorem store_outside_agree: - forall (chunk: memory_chunk) (m1 m2 m2': mem) - (b: block) (lo hi: Z) - (b': block) (ofs: Z) (v: val), - block_agree b lo hi m1 m2 -> - b <> b' \/ ofs + size_chunk chunk <= lo \/ hi <= ofs -> - store chunk m2 b' ofs v = Some m2' -> - block_agree b lo hi m1 m2'. -Proof. - intros. - generalize (store_inv _ _ _ _ _ _ H1). - intros [I [J [K [L [x M]]]]]. - red. red. rewrite M; unfold update; - case (zeq b b'); intro; simpl. - subst b'. apply store_contentmap_outside_agree. - assumption. - elim H0; intro. tauto. exact H2. - apply H. -Qed. - -(** * Store extensions *) - -(** A store [m2] extends a store [m1] if [m2] can be obtained from [m1] - by increasing the sizes of the memory blocks of [m1] (decreasing - the low bounds, increasing the high bounds), while still keeping the - same contents for block offsets that are valid in [m1]. - This notion is used in the proof of semantic equivalence in - module [Machenv]. *) - -Definition block_contents_extends (c1 c2: block_contents) := - c2.(low) <= c1.(low) /\ c1.(high) <= c2.(high) /\ - contentmap_agree c1.(low) c1.(high) c1.(contents) c2.(contents). - -Definition extends (m1 m2: mem) := - m1.(nextblock) = m2.(nextblock) /\ - forall (b: block), - b < m1.(nextblock) -> - block_contents_extends (m1.(blocks) b) (m2.(blocks) b). - -Theorem extends_refl: - forall (m: mem), extends m m. -Proof. - intro. red. split. - reflexivity. - intros. red. - split. omega. split. omega. - red. intros. reflexivity. -Qed. - -Theorem alloc_extends: - forall (m1 m2 m1' m2': mem) (lo1 hi1 lo2 hi2: Z) (b1 b2: block), - extends m1 m2 -> - lo2 <= lo1 -> hi1 <= hi2 -> - alloc m1 lo1 hi1 = (m1', b1) -> - alloc m2 lo2 hi2 = (m2', b2) -> - b1 = b2 /\ extends m1' m2'. -Proof. - unfold extends, alloc; intros. - injection H2; intros; subst m1'; subst b1. - injection H3; intros; subst m2'; subst b2. - simpl. intuition. - rewrite <- H4. case (zeq b (nextblock m1)); intro. - subst b. repeat rewrite update_s. - red; simpl. intuition. - red; intros; reflexivity. - repeat rewrite update_o. apply H5. omega. - auto. auto. -Qed. - -Theorem free_extends: - forall (m1 m2: mem) (b: block), - extends m1 m2 -> - extends (free m1 b) (free m2 b). -Proof. - unfold extends, free; intros. - simpl. intuition. - red; intros; simpl. - case (zeq b0 b); intro. - subst b0; repeat (rewrite update_s). - simpl. split. omega. split. omega. - red. intros. omegaContradiction. - repeat (rewrite update_o). - exact (H1 b0 H). - auto. auto. -Qed. - -Theorem load_extends: - forall (chunk: memory_chunk) (m1 m2: mem) (b: block) (ofs: Z) (v: val), - extends m1 m2 -> - load chunk m1 b ofs = Some v -> - load chunk m2 b ofs = Some v. -Proof. - intros sz m1 m2 b ofs v (X, Y) L. - generalize (load_inv _ _ _ _ _ L). - intros (A, (B, (C, D))). - unfold load. rewrite <- X. rewrite zlt_true; auto. - generalize (Y b A); intros [M [P Q]]. - rewrite in_bounds_holds. - rewrite <- D. - decEq. decEq. - apply load_contentmap_agree with - (lo := low((blocks m1) b)) - (hi := high((blocks m1) b)); auto. - omega. omega. -Qed. - -Theorem store_within_extends: - forall (chunk: memory_chunk) (m1 m2 m1': mem) (b: block) (ofs: Z) (v: val), - extends m1 m2 -> - store chunk m1 b ofs v = Some m1' -> - exists m2', store chunk m2 b ofs v = Some m2' /\ extends m1' m2'. -Proof. - intros sz m1 m2 m1' b ofs v (X, Y) STORE. - generalize (store_inv _ _ _ _ _ _ STORE). - intros (A, (B, (C, (D, (x, E))))). - generalize (Y b A); intros [M [P Q]]. - unfold store. rewrite <- X. rewrite zlt_true; auto. - case (in_bounds sz ofs (blocks m2 b)); intro. - exists (unchecked_store sz m2 b ofs v a). - split. auto. - split. - unfold unchecked_store; simpl. congruence. - intros b' LT. - unfold block_contents_extends, unchecked_store, block_contents_agree. - rewrite E; unfold update; simpl. - case (zeq b' b); intro; simpl. - subst b'. split. omega. split. omega. - apply store_contentmap_agree. auto. - apply Y. rewrite <- D. auto. - omegaContradiction. -Qed. - -Theorem store_outside_extends: - forall (chunk: memory_chunk) (m1 m2 m2': mem) (b: block) (ofs: Z) (v: val), - extends m1 m2 -> - ofs + size_chunk chunk <= low_bound m1 b \/ high_bound m1 b <= ofs -> - store chunk m2 b ofs v = Some m2' -> - extends m1 m2'. -Proof. - intros sz m1 m2 m2' b ofs v (X, Y) BOUNDS STORE. - generalize (store_inv _ _ _ _ _ _ STORE). - intros (A, (B, (C, (D, (x, E))))). - split. - congruence. - intros b' LT. - rewrite E; unfold update; case (zeq b' b); intro. - subst b'. generalize (Y b LT). intros [M [P Q]]. - red; simpl. split. omega. split. omega. - apply store_contentmap_outside_agree. - assumption. exact BOUNDS. - auto. -Qed. - -(** * Memory extensionality properties *) - -Lemma block_contents_exten: - forall (c1 c2: block_contents), - c1.(low) = c2.(low) -> - c1.(high) = c2.(high) -> - block_contents_agree c1.(low) c1.(high) c1 c2 -> - c1 = c2. -Proof. - intros. caseEq c1; intros lo1 hi1 m1 UO1 EQ1. subst c1. - caseEq c2; intros lo2 hi2 m2 UO2 EQ2. subst c2. - simpl in *. subst lo2 hi2. - assert (m1 = m2). - unfold contentmap. apply extensionality. intro. - case (zlt x lo1); intro. - rewrite UO1. rewrite UO2. auto. tauto. tauto. - case (zlt x hi1); intro. - apply H1. omega. - rewrite UO1. rewrite UO2. auto. tauto. tauto. - subst m2. - assert (UO1 = UO2). - apply proof_irrelevance. - subst UO2. reflexivity. -Qed. - -Theorem mem_exten: - forall m1 m2, - m1.(nextblock) = m2.(nextblock) -> - (forall b, m1.(blocks) b = m2.(blocks) b) -> - m1 = m2. -Proof. - intros. destruct m1 as [map1 nb1 nextpos1]. destruct m2 as [map2 nb2 nextpos2]. - simpl in *. subst nb2. - assert (map1 = map2). apply extensionality. assumption. - assert (nextpos1 = nextpos2). apply proof_irrelevance. - congruence. -Qed. - -(** * Store injections *) - -(** A memory injection [f] is a function from addresses to either [None] - or [Some] of an address and an offset. It defines a correspondence - between the blocks of two memory states [m1] and [m2]: -- if [f b = None], the block [b] of [m1] has no equivalent in [m2]; -- if [f b = Some(b', ofs)], the block [b] of [m2] corresponds to - a sub-block at offset [ofs] of the block [b'] in [m2]. -*) - -Definition meminj := (block -> option (block * Z))%type. - -Section MEM_INJECT. - -Variable f: meminj. - -(** A memory injection defines a relation between values that is the - identity relation, except for pointer values which are shifted - as prescribed by the memory injection. *) - -Inductive val_inject: val -> val -> Prop := - | val_inject_int: - forall i, val_inject (Vint i) (Vint i) - | val_inject_float: - forall f, val_inject (Vfloat f) (Vfloat f) - | val_inject_ptr: - forall b1 ofs1 b2 ofs2 x, - f b1 = Some (b2, x) -> - ofs2 = Int.add ofs1 (Int.repr x) -> - val_inject (Vptr b1 ofs1) (Vptr b2 ofs2) - | val_inject_undef: forall v, - val_inject Vundef v. - -Hint Resolve val_inject_int val_inject_float val_inject_ptr -val_inject_undef. - -Inductive val_list_inject: list val -> list val-> Prop:= - | val_nil_inject : - val_list_inject nil nil - | val_cons_inject : forall v v' vl vl' , - val_inject v v' -> val_list_inject vl vl'-> - val_list_inject (v::vl) (v':: vl'). - -Inductive val_content_inject: memory_size -> val -> val -> Prop := - | val_content_inject_base: - forall sz v1 v2, - val_inject v1 v2 -> - val_content_inject sz v1 v2 - | val_content_inject_8: - forall n1 n2, - Int.cast8unsigned n1 = Int.cast8unsigned n2 -> - val_content_inject Size8 (Vint n1) (Vint n2) - | val_content_inject_16: - forall n1 n2, - Int.cast16unsigned n1 = Int.cast16unsigned n2 -> - val_content_inject Size16 (Vint n1) (Vint n2) - | val_content_inject_32: - forall f1 f2, - Float.singleoffloat f1 = Float.singleoffloat f2 -> - val_content_inject Size32 (Vfloat f1) (Vfloat f2). - -Hint Resolve val_content_inject_base. - -Inductive content_inject: content -> content -> Prop := - | content_inject_undef: - forall c, - content_inject Undef c - | content_inject_datum8: - forall v1 v2, - val_content_inject Size8 v1 v2 -> - content_inject (Datum8 v1) (Datum8 v2) - | content_inject_datum16: - forall v1 v2, - val_content_inject Size16 v1 v2 -> - content_inject (Datum16 v1) (Datum16 v2) - | content_inject_datum32: - forall v1 v2, - val_content_inject Size32 v1 v2 -> - content_inject (Datum32 v1) (Datum32 v2) - | content_inject_datum64: - forall v1 v2, - val_content_inject Size64 v1 v2 -> - content_inject (Datum64 v1) (Datum64 v2) - | content_inject_cont: - content_inject Cont Cont. - -Hint Resolve content_inject_undef content_inject_datum8 -content_inject_datum16 content_inject_datum32 content_inject_datum64 -content_inject_cont. - -Definition contentmap_inject (c1 c2: contentmap) (lo hi delta: Z) : Prop := - forall x, lo <= x < hi -> - content_inject (c1 x) (c2 (x + delta)). - -(** A block contents [c1] injects into another block content [c2] at - offset [delta] if the contents of [c1] at all valid offsets - correspond to the contents of [c2] at the same offsets shifted by [delta]. - Some additional conditions are imposed to guard against arithmetic - overflow in address computations. *) - -Record block_contents_inject (c1 c2: block_contents) (delta: Z) : Prop := - mk_block_contents_inject { - bci_range1: Int.min_signed <= delta <= Int.max_signed; - bci_range2: delta = 0 \/ - (Int.min_signed <= c2.(low) /\ c2.(high) <= Int.max_signed); - bci_lowhigh:forall x, c1.(low) <= x < c1.(high) -> - c2.(low) <= x+delta < c2.(high) ; - bci_contents: - contentmap_inject c1.(contents) c2.(contents) c1.(low) c1.(high) delta - }. - -(** A memory state [m1] injects into another memory state [m2] via the - memory injection [f] if the following conditions hold: -- unallocated blocks in [m1] must be mapped to [None] by [f]; -- if [f b = Some(b', delta)], [b'] must be valid in [m2] and - the contents of [b] in [m1] must inject into the contents of [b'] in [m2] - with offset [delta]; -- distinct blocks in [m1] cannot be mapped to overlapping sub-blocks in [m2]. -*) - -Record mem_inject (m1 m2: mem) : Prop := - mk_mem_inject { - mi_freeblocks: - forall b, b >= m1.(nextblock) -> f b = None; - mi_mappedblocks: - forall b b' delta, - f b = Some(b', delta) -> - b' < m2.(nextblock) /\ - block_contents_inject (m1.(blocks) b) - (m2.(blocks) b') - delta; - mi_no_overlap: - forall b1 b2 b1' b2' delta1 delta2, - b1 <> b2 -> - f b1 = Some (b1', delta1) -> - f b2 = Some (b2', delta2) -> - b1' <> b2' - \/ (forall x1 x2, low_bound m1 b1 <= x1 < high_bound m1 b1 -> - low_bound m1 b2 <= x2 < high_bound m1 b2 -> - x1+delta1 <> x2+delta2) - }. - -(** The following lemmas establish the absence of machine integer overflow - during address computations. *) - -Lemma size_mem_pos: forall sz, size_mem sz > 0. -Proof. destruct sz; simpl; omega. Qed. - -Lemma size_chunk_pos: forall chunk, size_chunk chunk > 0. -Proof. intros. unfold size_chunk. apply size_mem_pos. Qed. - -Lemma address_inject: - forall m1 m2 chunk b1 ofs1 b2 ofs2 x, - mem_inject m1 m2 -> - (m1.(blocks) b1).(low) <= Int.signed ofs1 -> - Int.signed ofs1 + size_chunk chunk <= (m1.(blocks) b1).(high) -> - f b1 = Some (b2, x) -> - ofs2 = Int.add ofs1 (Int.repr x) -> - Int.signed ofs2 = Int.signed ofs1 + x. -Proof. - intros. - generalize (size_chunk_pos chunk). intro. - generalize (mi_mappedblocks m1 m2 H _ _ _ H2). intros [C D]. - inversion D. - elim bci_range4 ; intro. - (**x=0**) - subst x . rewrite Zplus_0_r. rewrite Int.add_zero in H3. - subst ofs2 ; auto . - (**x<>0**) - rewrite H3. rewrite Int.add_signed. repeat rewrite Int.signed_repr. - auto. auto. - assert (low (blocks m1 b1) <= Int.signed ofs1 < high (blocks m1 b1)). - omega. - generalize (bci_lowhigh0 (Int.signed ofs1) H6). omega. - auto. -Qed. - -Lemma valid_pointer_inject_no_overflow: - forall m1 m2 b ofs b' x, - mem_inject m1 m2 -> - valid_pointer m1 b (Int.signed ofs) = true -> - f b = Some(b', x) -> - Int.min_signed <= Int.signed ofs + Int.signed (Int.repr x) <= Int.max_signed. -Proof. - intros. unfold valid_pointer in H0. - destruct (zlt b (nextblock m1)); try discriminate. - destruct (zle (low (blocks m1 b)) (Int.signed ofs)); try discriminate. - destruct (zlt (Int.signed ofs) (high (blocks m1 b))); try discriminate. - inversion H. generalize (mi_mappedblocks0 _ _ _ H1). - intros [A B]. inversion B. - elim bci_range4 ; intro. - (**x=0**) - rewrite Int.signed_repr ; auto. - subst x . rewrite Zplus_0_r. apply Int.signed_range . - (**x<>0**) - generalize (bci_lowhigh0 _ (conj z0 z1)). intro. - rewrite Int.signed_repr. omega. auto. -Qed. - -(** Relation between injections and loads. *) - -Lemma check_cont_inject: - forall c1 c2 lo hi delta, - contentmap_inject c1 c2 lo hi delta -> - forall n p, - lo <= p -> p + Z_of_nat n <= hi -> - check_cont n p c1 = true -> - check_cont n (p + delta) c2 = true. -Proof. - induction n. - intros. simpl. reflexivity. - simpl check_cont. rewrite inj_S. intros p H0 H1. - assert (lo <= p < hi). omega. - generalize (H p H2). intro. inversion H3; intros; try discriminate. - replace (p + delta + 1) with ((p + 1) + delta). - apply IHn. omega. omega. auto. - omega. -Qed. - -Hint Resolve check_cont_inject. - -Lemma getN_inject: - forall c1 c2 lo hi delta, - contentmap_inject c1 c2 lo hi delta -> - forall n p, - lo <= p -> p + Z_of_nat n < hi -> - content_inject (getN n p c1) (getN n (p + delta) c2). -Proof. - intros. simpl in H1. - assert (lo <= p < hi). omega. - unfold getN. - caseEq (check_cont n (p + 1) c1); intro. - replace (check_cont n (p + delta + 1) c2) with true. - apply H. assumption. - symmetry. replace (p + delta + 1) with ((p + 1) + delta). - eapply check_cont_inject; eauto. - omega. omega. omega. - constructor. -Qed. - -Hint Resolve getN_inject. - -Definition ztonat (z:Z): nat:= -match z with -| Z0 => O -| Zpos p => (nat_of_P p) -| Zneg p =>O -end. - -Lemma load_contents_inject: - forall sz c1 c2 lo hi delta p, - contentmap_inject c1 c2 lo hi delta -> - lo <= p -> p + size_mem sz <= hi -> - val_content_inject sz (load_contents sz c1 p) (load_contents sz c2 (p + delta)). -Proof. -intros. -assert (content_inject (getN (ztonat (size_mem sz)-1) p c1) -(getN (ztonat(size_mem sz)-1) (p + delta) c2)). -induction sz; assert (lo<= p< hi); simpl in H1; try omega; -apply getN_inject with lo hi; try assumption; simpl ; try omega. -induction sz ; simpl; inversion H2; auto. -Qed. - -Hint Resolve load_contents_inject. - -Lemma load_result_inject: - forall chunk v1 v2, - val_content_inject (mem_chunk chunk) v1 v2 -> - val_inject (Val.load_result chunk v1) (Val.load_result chunk v2). -Proof. -intros. -destruct chunk; simpl in H; inversion H; simpl; auto; -try (inversion H0; simpl; auto; fail). -replace (Int.cast8signed n2) with (Int.cast8signed n1). constructor. -apply Int.cast8_signed_equal_if_unsigned_equal; auto. -rewrite H0; constructor. -replace (Int.cast16signed n2) with (Int.cast16signed n1). constructor. -apply Int.cast16_signed_equal_if_unsigned_equal; auto. -rewrite H0; constructor. -inversion H0; simpl; auto. -apply val_inject_ptr with x; auto. -Qed. - -Lemma in_bounds_inject: - forall chunk c1 c2 delta p, - block_contents_inject c1 c2 delta -> - c1.(low) <= p /\ p + size_chunk chunk <= c1.(high) -> - c2.(low) <= p + delta /\ (p + delta) + size_chunk chunk <= c2.(high). -Proof. - intros. - inversion H. - generalize (size_chunk_pos chunk); intro. - assert (low c1 <= p + size_chunk chunk - 1 < high c1). - omega. - generalize (bci_lowhigh0 _ H2). intro. - assert (low c1 <= p < high c1). - omega. - generalize (bci_lowhigh0 _ H4). intro. - omega. -Qed. - -Lemma block_cont_val: -forall c1 c2 delta p sz, -block_contents_inject c1 c2 delta -> -c1.(low) <= p -> p + size_mem sz <= c1.(high) -> - val_content_inject sz (load_contents sz c1.(contents) p) - (load_contents sz c2.(contents) (p + delta)). -Proof. -intros. -inversion H . -apply load_contents_inject with c1.(low) c1.(high) ;assumption. -Qed. - -Lemma load_inject: - forall m1 m2 chunk b1 ofs b2 delta v1, - mem_inject m1 m2 -> - load chunk m1 b1 ofs = Some v1 -> - f b1 = Some (b2, delta) -> - exists v2, load chunk m2 b2 (ofs + delta) = Some v2 /\ val_inject v1 v2. -Proof. - intros. - generalize (load_inv _ _ _ _ _ H0). intros [A [B [C D]]]. - inversion H. - generalize (mi_mappedblocks0 _ _ _ H1). intros [U V]. - inversion V. - exists (Val.load_result chunk (load_contents (mem_chunk chunk) - (m2.(blocks) b2).(contents) (ofs+delta))). - split. - unfold load. - generalize (size_chunk_pos chunk); intro. - rewrite zlt_true. rewrite in_bounds_holds. auto. - assert (low (blocks m1 b1) <= ofs < high (blocks m1 b1)). - omega. - generalize (bci_lowhigh0 _ H3). tauto. - assert (low (blocks m1 b1) <= ofs + size_chunk chunk - 1 < high(blocks m1 b1)). - omega. - generalize (bci_lowhigh0 _ H3). omega. - assumption. - rewrite <- D. apply load_result_inject. - eapply load_contents_inject; eauto. -Qed. - -Lemma loadv_inject: - forall m1 m2 chunk a1 a2 v1, - mem_inject m1 m2 -> - loadv chunk m1 a1 = Some v1 -> - val_inject a1 a2 -> - exists v2, loadv chunk m2 a2 = Some v2 /\ val_inject v1 v2. -Proof. - intros. - induction H1 ; simpl in H0 ; try discriminate H0. - simpl. - replace (Int.signed ofs2) with (Int.signed ofs1 + x). - apply load_inject with m1 b1 ; assumption. - symmetry. generalize (load_inv _ _ _ _ _ H0). intros [A [B [C D]]]. - apply address_inject with m1 m2 chunk b1 b2; auto. -Qed. - -(** Relation between injections and stores. *) - -Lemma set_cont_inject: - forall c1 c2 lo hi delta, - contentmap_inject c1 c2 lo hi delta -> - forall n p, - lo <= p -> p + Z_of_nat n <= hi -> - contentmap_inject (set_cont n p c1) (set_cont n (p + delta) c2) lo hi delta. -Proof. -induction n. intros. simpl. assumption. -intros. simpl. unfold contentmap_inject. -intros p2 Hp2. unfold update. -case (zeq p2 p); intro. -subst p2. rewrite zeq_true. constructor. -rewrite zeq_false. replace (p + delta + 1) with ((p + 1) + delta). -apply IHn. omega. rewrite inj_S in H1. omega. assumption. -omega. omega. -Qed. - -Lemma setN_inject: - forall c1 c2 lo hi delta n p v1 v2, - contentmap_inject c1 c2 lo hi delta -> - content_inject v1 v2 -> - lo <= p -> p + Z_of_nat n < hi -> - contentmap_inject (setN n p v1 c1) (setN n (p + delta) v2 c2) - lo hi delta. -Proof. - intros. unfold setN. red; intros. - unfold update. - case (zeq x p); intro. - subst p. rewrite zeq_true. assumption. - rewrite zeq_false. - replace (p + delta + 1) with ((p + 1) + delta). - apply set_cont_inject with lo hi. - assumption. omega. omega. assumption. omega. - omega. -Qed. - -Lemma store_contents_inject: - forall c1 c2 lo hi delta sz p v1 v2, - contentmap_inject c1 c2 lo hi delta -> - val_content_inject sz v1 v2 -> - lo <= p -> p + size_mem sz <= hi -> - contentmap_inject (store_contents sz c1 p v1) - (store_contents sz c2 (p + delta) v2) - lo hi delta. -Proof. - intros. - destruct sz; simpl in *; apply setN_inject; auto; simpl; omega. -Qed. - -Lemma set_cont_outside1 : - forall n p m q , - (forall x , p <= x < p + Z_of_nat n -> x <> q) -> - (set_cont n p m) q = m q. -Proof. - induction n; intros; simpl. - auto. - rewrite inj_S in H. rewrite update_o. apply IHn. - intros. apply H. omega. - apply H. omega. -Qed. - -Lemma set_cont_outside_inject: - forall c1 c2 lo hi delta, - contentmap_inject c1 c2 lo hi delta -> - forall n p, - (forall x1 x2, p <= x1 < p + Z_of_nat n -> - lo <= x2 < hi -> - x1 <> x2 + delta) -> - contentmap_inject c1 (set_cont n p c2) lo hi delta. -Proof. - unfold contentmap_inject. intros. - rewrite set_cont_outside1. apply H. assumption. - intros. apply H0. auto. auto. -Qed. - -Lemma setN_outside_inject: - forall n c1 c2 lo hi delta p v, - contentmap_inject c1 c2 lo hi delta -> - (forall x1 x2, p <= x1 < p + Z_of_nat (S n) -> - lo <= x2 < hi -> - x1 <> x2 + delta) -> - contentmap_inject c1 (setN n p v c2) lo hi delta. -Proof. - intros. unfold setN. red; intros. rewrite update_o. - apply set_cont_outside_inject with lo hi. auto. - intros. apply H0. rewrite inj_S. omega. auto. auto. - apply H0. rewrite inj_S. omega. auto. -Qed. - -Lemma store_contents_outside_inject: - forall c1 c2 lo hi delta sz p v, - contentmap_inject c1 c2 lo hi delta -> - (forall x1 x2, p <= x1 < p + size_mem sz -> - lo <= x2 < hi -> - x1 <> x2 + delta)-> - contentmap_inject c1 (store_contents sz c2 p v) lo hi delta. -Proof. - intros c1 c2 lo hi delta sz. generalize (size_mem_pos sz) . intro . - induction sz ;intros ;simpl ; apply setN_outside_inject ; assumption . -Qed. - -Lemma store_mapped_inject_1: - forall chunk m1 b1 ofs v1 n1 m2 b2 delta v2, - mem_inject m1 m2 -> - store chunk m1 b1 ofs v1 = Some n1 -> - f b1 = Some (b2, delta) -> - val_content_inject (mem_chunk chunk) v1 v2 -> - exists n2, - store chunk m2 b2 (ofs + delta) v2 = Some n2 - /\ mem_inject n1 n2. -Proof. -intros. -generalize (size_chunk_pos chunk); intro SIZEPOS. -generalize (store_inv _ _ _ _ _ _ H0). -intros [A [B [C [D [P E]]]]]. -inversion H. -generalize (mi_mappedblocks0 _ _ _ H1). intros [U V]. -inversion V. -generalize (in_bounds_inject _ _ _ _ _ V (conj B C)). intro BND. -exists (unchecked_store chunk m2 b2 (ofs+delta) v2 BND). -split. unfold store. -rewrite zlt_true; auto. -case (in_bounds chunk (ofs + delta) (blocks m2 b2)); intro. -assert (a = BND). apply proof_irrelevance. congruence. -omegaContradiction. -constructor. -intros. apply mi_freeblocks0. rewrite <- D. assumption. -intros. generalize (mi_mappedblocks0 b b' delta0 H3). -intros [W Y]. split. simpl. auto. -rewrite E; simpl. unfold update. -(* Cas 1: memes blocs b = b1 b'= b2 *) -case (zeq b b1); intro. -subst b. assert (b'= b2). congruence. subst b'. -assert (delta0 = delta). congruence. subst delta0. -rewrite zeq_true. inversion Y. constructor; simpl; auto. -apply store_contents_inject; auto. -(* Cas 2: blocs differents dans m1 mais mappes sur le meme bloc de m2 *) -case (zeq b' b2); intro. -subst b'. -inversion Y. constructor; simpl; auto. -generalize (store_contents_outside_inject _ _ _ _ _ (mem_chunk chunk) - (ofs+delta) v2 bci_contents1). -intros. -apply H4. -elim (mi_no_overlap0 b b1 b2 b2 delta0 delta n H3 H1). -tauto. -unfold high_bound, low_bound. intros. -apply sym_not_equal. replace x1 with ((x1 - delta) + delta). -apply H5. assumption. unfold size_chunk in C. omega. omega. -(* Cas 3: blocs differents dans m1 et dans m2 *) -auto. - -unfold high_bound, low_bound. -rewrite E; simpl; intros. -unfold high_bound, low_bound in mi_no_overlap0. -unfold update. -case (zeq b0 b1). -intro. subst b1. simpl. -rewrite zeq_false; auto. -intro. case (zeq b3 b1); intro. -subst b3. simpl. auto. -auto. -Qed. - -Lemma store_mapped_inject: - forall chunk m1 b1 ofs v1 n1 m2 b2 delta v2, - mem_inject m1 m2 -> - store chunk m1 b1 ofs v1 = Some n1 -> - f b1 = Some (b2, delta) -> - val_inject v1 v2 -> - exists n2, - store chunk m2 b2 (ofs + delta) v2 = Some n2 - /\ mem_inject n1 n2. -Proof. - intros. eapply store_mapped_inject_1; eauto. -Qed. - -Lemma store_unmapped_inject: - forall chunk m1 b1 ofs v1 n1 m2, - mem_inject m1 m2 -> - store chunk m1 b1 ofs v1 = Some n1 -> - f b1 = None -> - mem_inject n1 m2. -Proof. -intros. -inversion H. -generalize (store_inv _ _ _ _ _ _ H0). -intros [A [B [C [D [P E]]]]]. -constructor. -rewrite D. assumption. -intros. elim (mi_mappedblocks0 _ _ _ H2); intros. -split. auto. -rewrite E; unfold update. -rewrite zeq_false. assumption. -congruence. -intros. -assert (forall b, low_bound n1 b = low_bound m1 b). - intros. unfold low_bound; rewrite E; unfold update. - case (zeq b b1); intros. subst b1; reflexivity. reflexivity. -assert (forall b, high_bound n1 b = high_bound m1 b). - intros. unfold high_bound; rewrite E; unfold update. - case (zeq b b1); intros. subst b1; reflexivity. reflexivity. -repeat rewrite H5. repeat rewrite H6. auto. -Qed. - -Lemma storev_mapped_inject_1: - forall chunk m1 a1 v1 n1 m2 a2 v2, - mem_inject m1 m2 -> - storev chunk m1 a1 v1 = Some n1 -> - val_inject a1 a2 -> - val_content_inject (mem_chunk chunk) v1 v2 -> - exists n2, - storev chunk m2 a2 v2 = Some n2 /\ mem_inject n1 n2. -Proof. - intros. destruct a1; simpl in H0; try discriminate. - inversion H1. subst b. - simpl. replace (Int.signed ofs2) with (Int.signed i + x). - eapply store_mapped_inject_1; eauto. - symmetry. generalize (store_inv _ _ _ _ _ _ H0). intros [A [B [C [D [P E]]]]]. - apply address_inject with m1 m2 chunk b1 b2; auto. -Qed. - -Lemma storev_mapped_inject: - forall chunk m1 a1 v1 n1 m2 a2 v2, - mem_inject m1 m2 -> - storev chunk m1 a1 v1 = Some n1 -> - val_inject a1 a2 -> - val_inject v1 v2 -> - exists n2, - storev chunk m2 a2 v2 = Some n2 /\ mem_inject n1 n2. -Proof. - intros. eapply storev_mapped_inject_1; eauto. -Qed. - -(** Relation between injections and [free] *) - -Lemma free_first_inject : - forall m1 m2 b, - mem_inject m1 m2 -> - mem_inject (free m1 b) m2. -Proof. - intros. inversion H. constructor . auto. - simpl. intros. - generalize (mi_mappedblocks0 b0 b' delta H0). - intros [A B] . split. assumption . unfold update. - case (zeq b0 b); intro. inversion B. constructor; simpl; auto. - intros. omega. - unfold contentmap_inject. - intros. omegaContradiction. - auto. intros. - unfold free . unfold low_bound , high_bound. simpl. unfold update. - case (zeq b1 b);intro. simpl. - right. intros. omegaContradiction. - case (zeq b2 b);intro. simpl. - right . intros. omegaContradiction. - unfold low_bound, high_bound in mi_no_overlap0. auto. -Qed. - -Lemma free_first_list_inject : - forall l m1 m2, - mem_inject m1 m2 -> - mem_inject (free_list m1 l) m2. -Proof. - induction l; simpl; intros. - auto. - apply free_first_inject. auto. -Qed. - -Lemma free_snd_inject: - forall m1 m2 b, - (forall b1 delta, f b1 = Some(b, delta) -> - low_bound m1 b1 >= high_bound m1 b1) -> - mem_inject m1 m2 -> - mem_inject m1 (free m2 b). -Proof. - intros. inversion H0. constructor. auto. - simpl. intros. - generalize (mi_mappedblocks0 b0 b' delta H1). - intros [A B] . split. assumption . - inversion B. unfold update. - case (zeq b' b); intro. - subst b'. generalize (H _ _ H1). unfold low_bound, high_bound; intro. - constructor. auto. elim bci_range4 ; intro. - (**delta=0**) - left ; auto . - (** delta<>0**) - right . - simpl. compute. split; intro; congruence. - intros. omegaContradiction. - red; intros. omegaContradiction. - auto. - auto. -Qed. - -Lemma bounds_free_block: - forall m1 b m1' , free m1 b = m1' -> - low_bound m1' b >= high_bound m1' b. -Proof. - intros. unfold free in H. - rewrite<- H . unfold low_bound , high_bound . - simpl . rewrite update_s. simpl. omega. -Qed. - -Lemma free_empty_bounds: - forall l m1 , - let m1' := free_list m1 l in - (forall b, In b l -> low_bound m1' b >= high_bound m1' b). -Proof. - induction l . intros . inversion H. - intros. - generalize (in_inv H). - intro . elim H0. intro . subst b. simpl in m1' . - generalize ( bounds_free_block - (fold_right (fun (b : block) (m : mem) => free m b) m1 l) a m1') . - intros. apply H1. auto . intros. simpl in m1'. unfold m1' . - unfold low_bound , high_bound . simpl . - unfold update; case (zeq b a); intro; simpl. - omega . - unfold low_bound , high_bound in IHl . auto. -Qed. - -Lemma free_inject: - forall m1 m2 l b, - (forall b1 delta, f b1 = Some(b, delta) -> In b1 l) -> - mem_inject m1 m2 -> - mem_inject (free_list m1 l) (free m2 b). -Proof. - intros. apply free_snd_inject. - intros. apply free_empty_bounds. apply H with delta. auto. - apply free_first_list_inject. auto. -Qed. - -Lemma contents_init_data_inject: - forall id, contentmap_inject (contents_init_data 0 id) (contents_init_data 0 id) 0 (size_init_data_list id) 0. -Proof. - intro id0. - set (sz0 := size_init_data_list id0). - assert (forall id pos, - 0 <= pos -> pos + size_init_data_list id <= sz0 -> - contentmap_inject (contents_init_data pos id) (contents_init_data pos id) 0 sz0 0). - induction id; simpl; intros. - red; intros; constructor. - assert (forall n dt, - size_init_data a = 1 + Z_of_nat n -> - content_inject dt dt -> - contentmap_inject (setN n pos dt (contents_init_data (pos + size_init_data a) id)) - (setN n pos dt (contents_init_data (pos + size_init_data a) id)) - 0 sz0 0). - intros. set (pos' := pos) in |- * at 3. replace pos' with (pos + 0). - apply setN_inject. apply IHid. omega. omega. auto. auto. - generalize (size_init_data_list_pos id). omega. unfold pos'; omega. - destruct a; - try (apply H1; [reflexivity|repeat constructor]). - apply IHid. generalize (Zmax2 z 0). omega. simpl in H0; omega. - - apply H. omega. unfold sz0. omega. -Qed. - -End MEM_INJECT. - -Hint Resolve val_inject_int val_inject_float val_inject_ptr val_inject_undef. -Hint Resolve val_nil_inject val_cons_inject. - -(** Monotonicity properties of memory injections. *) - -Definition inject_incr (f1 f2: meminj) : Prop := - forall b, f1 b = f2 b \/ f1 b = None. - -Lemma inject_incr_refl : - forall f , inject_incr f f . -Proof. unfold inject_incr . intros. left . auto . Qed. - -Lemma inject_incr_trans : - forall f1 f2 f3 , - inject_incr f1 f2 -> inject_incr f2 f3 -> inject_incr f1 f3 . -Proof . - unfold inject_incr . intros . - generalize (H b) . intro . generalize (H0 b) . intro . - elim H1 ; elim H2 ; intros. - rewrite H3 in H4 . left . auto . - rewrite H3 in H4 . right . auto . - right ; auto . - right ; auto . -Qed. - -Lemma val_inject_incr: - forall f1 f2 v v', - inject_incr f1 f2 -> - val_inject f1 v v' -> - val_inject f2 v v'. -Proof. - intros. inversion H0. - constructor. - constructor. - elim (H b1); intro. - apply val_inject_ptr with x. congruence. auto. - congruence. - constructor. -Qed. - -Lemma val_list_inject_incr: - forall f1 f2 vl vl' , - inject_incr f1 f2 -> val_list_inject f1 vl vl' -> - val_list_inject f2 vl vl'. -Proof. - induction vl ; intros; inversion H0. auto . - subst a . generalize (val_inject_incr f1 f2 v v' H H3) . - intro . auto . -Qed. - -Hint Resolve inject_incr_refl val_inject_incr val_list_inject_incr. - -Section MEM_INJECT_INCR. - -Variable f1 f2: meminj. -Hypothesis INCR: inject_incr f1 f2. - -Lemma val_content_inject_incr: - forall chunk v v', - val_content_inject f1 chunk v v' -> - val_content_inject f2 chunk v v'. -Proof. - intros. inversion H. - apply val_content_inject_base. eapply val_inject_incr; eauto. - apply val_content_inject_8; auto. - apply val_content_inject_16; auto. - apply val_content_inject_32; auto. -Qed. - -Lemma content_inject_incr: - forall c c', content_inject f1 c c' -> content_inject f2 c c'. -Proof. - induction 1; constructor; eapply val_content_inject_incr; eauto. -Qed. - -Lemma contentmap_inject_incr: - forall c c' lo hi delta, - contentmap_inject f1 c c' lo hi delta -> - contentmap_inject f2 c c' lo hi delta. -Proof. - unfold contentmap_inject; intros. - apply content_inject_incr. auto. -Qed. - -Lemma block_contents_inject_incr: - forall c c' delta, - block_contents_inject f1 c c' delta -> - block_contents_inject f2 c c' delta. -Proof. - intros. inversion H. constructor; auto. - apply contentmap_inject_incr; auto. -Qed. - -End MEM_INJECT_INCR. - -(** Properties of injections and allocations. *) - -Definition extend_inject - (b: block) (x: option (block * Z)) (f: meminj) : meminj := - fun b' => if eq_block b' b then x else f b'. - -Lemma extend_inject_incr: - forall f b x, - f b = None -> - inject_incr f (extend_inject b x f). -Proof. - intros; red; intros. unfold extend_inject. - case (eq_block b0 b); intro. subst b0. right; auto. left; auto. -Qed. - -Lemma alloc_right_inject: - forall f m1 m2 lo hi m2' b, - mem_inject f m1 m2 -> - alloc m2 lo hi = (m2', b) -> - mem_inject f m1 m2'. -Proof. - intros. unfold alloc in H0. injection H0; intros A B; clear H0. - inversion H. - constructor. - assumption. - intros. generalize (mi_mappedblocks0 _ _ _ H0). intros [C D]. - rewrite <- B; simpl. split. omega. - rewrite update_o. auto. omega. - assumption. -Qed. - -Lemma alloc_unmapped_inject: - forall f m1 m2 lo hi m1' b, - mem_inject f m1 m2 -> - alloc m1 lo hi = (m1', b) -> - mem_inject (extend_inject b None f) m1' m2 /\ - inject_incr f (extend_inject b None f). -Proof. - intros. unfold alloc in H0. injection H0; intros; clear H0. - inversion H. - assert (inject_incr f (extend_inject b None f)). - apply extend_inject_incr. apply mi_freeblocks0. rewrite H1. omega. - split; auto. - constructor. - rewrite <- H2; simpl; intros. unfold extend_inject. - case (eq_block b0 b); intro. auto. apply mi_freeblocks0. omega. - intros until delta. unfold extend_inject at 1. case (eq_block b0 b); intro. - intros; discriminate. - intros. rewrite <- H2; simpl. rewrite H1. - rewrite update_o; auto. - elim (mi_mappedblocks0 _ _ _ H3). intros A B. - split. auto. apply block_contents_inject_incr with f. auto. auto. - intros until delta2. unfold extend_inject. - case (eq_block b1 b); intro. congruence. - case (eq_block b2 b); intro. congruence. - rewrite <- H2. unfold low_bound, high_bound; simpl. rewrite H1. - repeat rewrite update_o; auto. - exact (mi_no_overlap0 b1 b2 b1' b2' delta1 delta2). -Qed. - -Lemma alloc_mapped_inject: - forall f m1 m2 lo hi m1' b b' ofs, - mem_inject f m1 m2 -> - alloc m1 lo hi = (m1', b) -> - valid_block m2 b' -> - Int.min_signed <= ofs <= Int.max_signed -> - Int.min_signed <= low_bound m2 b' -> - high_bound m2 b' <= Int.max_signed -> - low_bound m2 b' <= lo + ofs -> - hi + ofs <= high_bound m2 b' -> - (forall b0 ofs0, - f b0 = Some (b', ofs0) -> - high_bound m1 b0 + ofs0 <= lo + ofs \/ - hi + ofs <= low_bound m1 b0 + ofs0) -> - mem_inject (extend_inject b (Some (b', ofs)) f) m1' m2 /\ - inject_incr f (extend_inject b (Some (b', ofs)) f). -Proof. - intros. - generalize (fun b' => low_bound_alloc _ _ b' _ _ _ H0). - intro LOW. - generalize (fun b' => high_bound_alloc _ _ b' _ _ _ H0). - intro HIGH. - unfold alloc in H0. injection H0; intros A B; clear H0. - inversion H. - (* inject_incr *) - assert (inject_incr f (extend_inject b (Some (b', ofs)) f)). - apply extend_inject_incr. apply mi_freeblocks0. rewrite A. omega. - split; auto. - constructor. - (* mi_freeblocks *) - rewrite <- B; simpl; intros. unfold extend_inject. - case (eq_block b0 b); intro. unfold block in *. omegaContradiction. - apply mi_freeblocks0. omega. - (* mi_mappedblocks *) - intros until delta. unfold extend_inject at 1. - case (eq_block b0 b); intro. - intros. subst b0. inversion H8. subst b'0; subst delta. - split. assumption. - rewrite <- B; simpl. rewrite A. rewrite update_s. - constructor; auto. - unfold empty_block. simpl. intros. unfold low_bound in H5. unfold high_bound in H6. omega. - simpl. red; intros. constructor. - intros. - generalize (mi_mappedblocks0 _ _ _ H8). intros [C D]. - split. auto. - rewrite <- B; simpl; rewrite A; rewrite update_o; auto. - apply block_contents_inject_incr with f; auto. - (* no overlap *) - intros until delta2. unfold extend_inject. - repeat rewrite LOW. repeat rewrite HIGH. unfold eq_block. - case (zeq b1 b); case (zeq b2 b); intros. - congruence. - inversion H9. subst b1'; subst delta1. - case (eq_block b' b2'); intro. - subst b2'. generalize (H7 _ _ H10). intro. - right. intros. omega. left. auto. - inversion H10. subst b2'; subst delta2. - case (eq_block b' b1'); intro. - subst b1'. generalize (H7 _ _ H9). intro. - right. intros. omega. left. auto. - apply mi_no_overlap0; auto. -Qed. - -Lemma alloc_parallel_inject: - forall f m1 m2 lo hi m1' m2' b1 b2, - mem_inject f m1 m2 -> - alloc m1 lo hi = (m1', b1) -> - alloc m2 lo hi = (m2', b2) -> - Int.min_signed <= lo -> hi <= Int.max_signed -> - mem_inject (extend_inject b1 (Some(b2,0)) f) m1' m2' /\ - inject_incr f (extend_inject b1 (Some(b2,0)) f). -Proof. - intros. - generalize (low_bound_alloc _ _ b2 _ _ _ H1). rewrite zeq_true; intro LOW. - generalize (high_bound_alloc _ _ b2 _ _ _ H1). rewrite zeq_true; intro HIGH. - eapply alloc_mapped_inject; eauto. - eapply alloc_right_inject; eauto. - eapply valid_new_block; eauto. - compute. intuition congruence. - rewrite LOW; auto. - rewrite HIGH; auto. - rewrite LOW; omega. - rewrite HIGH; omega. - intros. elim (mi_mappedblocks _ _ _ H _ _ _ H4); intros. - injection H1; intros. rewrite H7 in H5. omegaContradiction. -Qed. diff --git a/backend/Values.v b/backend/Values.v deleted file mode 100644 index aa59e045..00000000 --- a/backend/Values.v +++ /dev/null @@ -1,888 +0,0 @@ -(** This module defines the type of values that is used in the dynamic - semantics of all our intermediate languages. *) - -Require Import Coqlib. -Require Import AST. -Require Import Integers. -Require Import Floats. - -Definition block : Set := Z. -Definition eq_block := zeq. - -(** A value is either: -- a machine integer; -- a floating-point number; -- a pointer: a pair of a memory address and an integer offset with respect - to this address; -- the [Vundef] value denoting an arbitrary bit pattern, such as the - value of an uninitialized variable. -*) - -Inductive val: Set := - | Vundef: val - | Vint: int -> val - | Vfloat: float -> val - | Vptr: block -> int -> val. - -Definition Vzero: val := Vint Int.zero. -Definition Vone: val := Vint Int.one. -Definition Vmone: val := Vint Int.mone. - -Definition Vtrue: val := Vint Int.one. -Definition Vfalse: val := Vint Int.zero. - -(** The module [Val] defines a number of arithmetic and logical operations - over type [val]. Most of these operations are straightforward extensions - of the corresponding integer or floating-point operations. *) - -Module Val. - -Definition of_bool (b: bool): val := if b then Vtrue else Vfalse. - -Definition has_type (v: val) (t: typ) : Prop := - match v, t with - | Vundef, _ => True - | Vint _, Tint => True - | Vfloat _, Tfloat => True - | Vptr _ _, Tint => True - | _, _ => False - end. - -Fixpoint has_type_list (vl: list val) (tl: list typ) {struct vl} : Prop := - match vl, tl with - | nil, nil => True - | v1 :: vs, t1 :: ts => has_type v1 t1 /\ has_type_list vs ts - | _, _ => False - end. - -(** Truth values. Pointers and non-zero integers are treated as [True]. - The integer 0 (also used to represent the null pointer) is [False]. - [Vundef] and floats are neither true nor false. *) - -Definition is_true (v: val) : Prop := - match v with - | Vint n => n <> Int.zero - | Vptr b ofs => True - | _ => False - end. - -Definition is_false (v: val) : Prop := - match v with - | Vint n => n = Int.zero - | _ => False - end. - -Inductive bool_of_val: val -> bool -> Prop := - | bool_of_val_int_true: - forall n, n <> Int.zero -> bool_of_val (Vint n) true - | bool_of_val_int_false: - bool_of_val (Vint Int.zero) false - | bool_of_val_ptr: - forall b ofs, bool_of_val (Vptr b ofs) true. - -Definition neg (v: val) : val := - match v with - | Vint n => Vint (Int.neg n) - | _ => Vundef - end. - -Definition negf (v: val) : val := - match v with - | Vfloat f => Vfloat (Float.neg f) - | _ => Vundef - end. - -Definition absf (v: val) : val := - match v with - | Vfloat f => Vfloat (Float.abs f) - | _ => Vundef - end. - -Definition intoffloat (v: val) : val := - match v with - | Vfloat f => Vint (Float.intoffloat f) - | _ => Vundef - end. - -Definition floatofint (v: val) : val := - match v with - | Vint n => Vfloat (Float.floatofint n) - | _ => Vundef - end. - -Definition floatofintu (v: val) : val := - match v with - | Vint n => Vfloat (Float.floatofintu n) - | _ => Vundef - end. - -Definition notint (v: val) : val := - match v with - | Vint n => Vint (Int.xor n Int.mone) - | _ => Vundef - end. - -Definition notbool (v: val) : val := - match v with - | Vint n => of_bool (Int.eq n Int.zero) - | Vptr b ofs => Vfalse - | _ => Vundef - end. - -Definition cast8signed (v: val) : val := - match v with - | Vint n => Vint(Int.cast8signed n) - | _ => Vundef - end. - -Definition cast8unsigned (v: val) : val := - match v with - | Vint n => Vint(Int.cast8unsigned n) - | _ => Vundef - end. - -Definition cast16signed (v: val) : val := - match v with - | Vint n => Vint(Int.cast16signed n) - | _ => Vundef - end. - -Definition cast16unsigned (v: val) : val := - match v with - | Vint n => Vint(Int.cast16unsigned n) - | _ => Vundef - end. - -Definition singleoffloat (v: val) : val := - match v with - | Vfloat f => Vfloat(Float.singleoffloat f) - | _ => Vundef - end. - -Definition add (v1 v2: val): val := - match v1, v2 with - | Vint n1, Vint n2 => Vint(Int.add n1 n2) - | Vptr b1 ofs1, Vint n2 => Vptr b1 (Int.add ofs1 n2) - | Vint n1, Vptr b2 ofs2 => Vptr b2 (Int.add ofs2 n1) - | _, _ => Vundef - end. - -Definition sub (v1 v2: val): val := - match v1, v2 with - | Vint n1, Vint n2 => Vint(Int.sub n1 n2) - | Vptr b1 ofs1, Vint n2 => Vptr b1 (Int.sub ofs1 n2) - | Vptr b1 ofs1, Vptr b2 ofs2 => - if zeq b1 b2 then Vint(Int.sub ofs1 ofs2) else Vundef - | _, _ => Vundef - end. - -Definition mul (v1 v2: val): val := - match v1, v2 with - | Vint n1, Vint n2 => Vint(Int.mul n1 n2) - | _, _ => Vundef - end. - -Definition divs (v1 v2: val): val := - match v1, v2 with - | Vint n1, Vint n2 => - if Int.eq n2 Int.zero then Vundef else Vint(Int.divs n1 n2) - | _, _ => Vundef - end. - -Definition mods (v1 v2: val): val := - match v1, v2 with - | Vint n1, Vint n2 => - if Int.eq n2 Int.zero then Vundef else Vint(Int.mods n1 n2) - | _, _ => Vundef - end. - -Definition divu (v1 v2: val): val := - match v1, v2 with - | Vint n1, Vint n2 => - if Int.eq n2 Int.zero then Vundef else Vint(Int.divu n1 n2) - | _, _ => Vundef - end. - -Definition modu (v1 v2: val): val := - match v1, v2 with - | Vint n1, Vint n2 => - if Int.eq n2 Int.zero then Vundef else Vint(Int.modu n1 n2) - | _, _ => Vundef - end. - -Definition and (v1 v2: val): val := - match v1, v2 with - | Vint n1, Vint n2 => Vint(Int.and n1 n2) - | _, _ => Vundef - end. - -Definition or (v1 v2: val): val := - match v1, v2 with - | Vint n1, Vint n2 => Vint(Int.or n1 n2) - | _, _ => Vundef - end. - -Definition xor (v1 v2: val): val := - match v1, v2 with - | Vint n1, Vint n2 => Vint(Int.xor n1 n2) - | _, _ => Vundef - end. - -Definition shl (v1 v2: val): val := - match v1, v2 with - | Vint n1, Vint n2 => - if Int.ltu n2 (Int.repr 32) - then Vint(Int.shl n1 n2) - else Vundef - | _, _ => Vundef - end. - -Definition shr (v1 v2: val): val := - match v1, v2 with - | Vint n1, Vint n2 => - if Int.ltu n2 (Int.repr 32) - then Vint(Int.shr n1 n2) - else Vundef - | _, _ => Vundef - end. - -Definition shr_carry (v1 v2: val): val := - match v1, v2 with - | Vint n1, Vint n2 => - if Int.ltu n2 (Int.repr 32) - then Vint(Int.shr_carry n1 n2) - else Vundef - | _, _ => Vundef - end. - -Definition shrx (v1 v2: val): val := - match v1, v2 with - | Vint n1, Vint n2 => - if Int.ltu n2 (Int.repr 32) - then Vint(Int.shrx n1 n2) - else Vundef - | _, _ => Vundef - end. - -Definition shru (v1 v2: val): val := - match v1, v2 with - | Vint n1, Vint n2 => - if Int.ltu n2 (Int.repr 32) - then Vint(Int.shru n1 n2) - else Vundef - | _, _ => Vundef - end. - -Definition rolm (v: val) (amount mask: int): val := - match v with - | Vint n => Vint(Int.rolm n amount mask) - | _ => Vundef - end. - -Definition addf (v1 v2: val): val := - match v1, v2 with - | Vfloat f1, Vfloat f2 => Vfloat(Float.add f1 f2) - | _, _ => Vundef - end. - -Definition subf (v1 v2: val): val := - match v1, v2 with - | Vfloat f1, Vfloat f2 => Vfloat(Float.sub f1 f2) - | _, _ => Vundef - end. - -Definition mulf (v1 v2: val): val := - match v1, v2 with - | Vfloat f1, Vfloat f2 => Vfloat(Float.mul f1 f2) - | _, _ => Vundef - end. - -Definition divf (v1 v2: val): val := - match v1, v2 with - | Vfloat f1, Vfloat f2 => Vfloat(Float.div f1 f2) - | _, _ => Vundef - end. - -Definition cmp_mismatch (c: comparison): val := - match c with - | Ceq => Vfalse - | Cne => Vtrue - | _ => Vundef - end. - -Definition cmp (c: comparison) (v1 v2: val): val := - match v1, v2 with - | Vint n1, Vint n2 => of_bool (Int.cmp c n1 n2) - | Vint n1, Vptr b2 ofs2 => - if Int.eq n1 Int.zero then cmp_mismatch c else Vundef - | Vptr b1 ofs1, Vptr b2 ofs2 => - if zeq b1 b2 - then of_bool (Int.cmp c ofs1 ofs2) - else cmp_mismatch c - | Vptr b1 ofs1, Vint n2 => - if Int.eq n2 Int.zero then cmp_mismatch c else Vundef - | _, _ => Vundef - end. - -Definition cmpu (c: comparison) (v1 v2: val): val := - match v1, v2 with - | Vint n1, Vint n2 => - of_bool (Int.cmpu c n1 n2) - | Vint n1, Vptr b2 ofs2 => - if Int.eq n1 Int.zero then cmp_mismatch c else Vundef - | Vptr b1 ofs1, Vptr b2 ofs2 => - if zeq b1 b2 - then of_bool (Int.cmpu c ofs1 ofs2) - else cmp_mismatch c - | Vptr b1 ofs1, Vint n2 => - if Int.eq n2 Int.zero then cmp_mismatch c else Vundef - | _, _ => Vundef - end. - -Definition cmpf (c: comparison) (v1 v2: val): val := - match v1, v2 with - | Vfloat f1, Vfloat f2 => of_bool (Float.cmp c f1 f2) - | _, _ => Vundef - end. - -(** [load_result] is used in the memory model (library [Mem]) - to post-process the results of a memory read. For instance, - consider storing the integer value [0xFFF] on 1 byte at a - given address, and reading it back. If it is read back with - chunk [Mint8unsigned], zero-extension must be performed, resulting - in [0xFF]. If it is read back as a [Mint8signed], sign-extension - is performed and [0xFFFFFFFF] is returned. Type mismatches - (e.g. reading back a float as a [Mint32]) read back as [Vundef]. *) - -Definition load_result (chunk: memory_chunk) (v: val) := - match chunk, v with - | Mint8signed, Vint n => Vint (Int.cast8signed n) - | Mint8unsigned, Vint n => Vint (Int.cast8unsigned n) - | Mint16signed, Vint n => Vint (Int.cast16signed n) - | Mint16unsigned, Vint n => Vint (Int.cast16unsigned n) - | Mint32, Vint n => Vint n - | Mint32, Vptr b ofs => Vptr b ofs - | Mfloat32, Vfloat f => Vfloat(Float.singleoffloat f) - | Mfloat64, Vfloat f => Vfloat f - | _, _ => Vundef - end. - -(** Theorems on arithmetic operations. *) - -Theorem cast8unsigned_and: - forall x, cast8unsigned x = and x (Vint(Int.repr 255)). -Proof. - destruct x; simpl; auto. decEq. apply Int.cast8unsigned_and. -Qed. - -Theorem cast16unsigned_and: - forall x, cast16unsigned x = and x (Vint(Int.repr 65535)). -Proof. - destruct x; simpl; auto. decEq. apply Int.cast16unsigned_and. -Qed. - -Theorem istrue_not_isfalse: - forall v, is_false v -> is_true (notbool v). -Proof. - destruct v; simpl; try contradiction. - intros. subst i. simpl. discriminate. -Qed. - -Theorem isfalse_not_istrue: - forall v, is_true v -> is_false (notbool v). -Proof. - destruct v; simpl; try contradiction. - intros. generalize (Int.eq_spec i Int.zero). - case (Int.eq i Int.zero); intro. - contradiction. simpl. auto. - auto. -Qed. - -Theorem bool_of_true_val: - forall v, is_true v -> bool_of_val v true. -Proof. - intro. destruct v; simpl; intros; try contradiction. - constructor; auto. constructor. -Qed. - -Theorem bool_of_true_val2: - forall v, bool_of_val v true -> is_true v. -Proof. - intros. inversion H; simpl; auto. -Qed. - -Theorem bool_of_true_val_inv: - forall v b, is_true v -> bool_of_val v b -> b = true. -Proof. - intros. inversion H0; subst v b; simpl in H; auto. -Qed. - -Theorem bool_of_false_val: - forall v, is_false v -> bool_of_val v false. -Proof. - intro. destruct v; simpl; intros; try contradiction. - subst i; constructor. -Qed. - -Theorem bool_of_false_val2: - forall v, bool_of_val v false -> is_false v. -Proof. - intros. inversion H; simpl; auto. -Qed. - -Theorem bool_of_false_val_inv: - forall v b, is_false v -> bool_of_val v b -> b = false. -Proof. - intros. inversion H0; subst v b; simpl in H. - congruence. auto. contradiction. -Qed. - -Theorem notbool_negb_1: - forall b, of_bool (negb b) = notbool (of_bool b). -Proof. - destruct b; reflexivity. -Qed. - -Theorem notbool_negb_2: - forall b, of_bool b = notbool (of_bool (negb b)). -Proof. - destruct b; reflexivity. -Qed. - -Theorem notbool_idem2: - forall b, notbool(notbool(of_bool b)) = of_bool b. -Proof. - destruct b; reflexivity. -Qed. - -Theorem notbool_idem3: - forall x, notbool(notbool(notbool x)) = notbool x. -Proof. - destruct x; simpl; auto. - case (Int.eq i Int.zero); reflexivity. -Qed. - -Theorem add_commut: forall x y, add x y = add y x. -Proof. - destruct x; destruct y; simpl; auto. - decEq. apply Int.add_commut. -Qed. - -Theorem add_assoc: forall x y z, add (add x y) z = add x (add y z). -Proof. - destruct x; destruct y; destruct z; simpl; auto. - rewrite Int.add_assoc; auto. - rewrite Int.add_assoc; auto. - decEq. decEq. apply Int.add_commut. - decEq. rewrite Int.add_commut. rewrite <- Int.add_assoc. - decEq. apply Int.add_commut. - decEq. rewrite Int.add_assoc. auto. -Qed. - -Theorem add_permut: forall x y z, add x (add y z) = add y (add x z). -Proof. - intros. rewrite (add_commut y z). rewrite <- add_assoc. apply add_commut. -Qed. - -Theorem add_permut_4: - forall x y z t, add (add x y) (add z t) = add (add x z) (add y t). -Proof. - intros. rewrite add_permut. rewrite add_assoc. - rewrite add_permut. symmetry. apply add_assoc. -Qed. - -Theorem neg_zero: neg Vzero = Vzero. -Proof. - reflexivity. -Qed. - -Theorem neg_add_distr: forall x y, neg(add x y) = add (neg x) (neg y). -Proof. - destruct x; destruct y; simpl; auto. decEq. apply Int.neg_add_distr. -Qed. - -Theorem sub_zero_r: forall x, sub Vzero x = neg x. -Proof. - destruct x; simpl; auto. -Qed. - -Theorem sub_add_opp: forall x y, sub x (Vint y) = add x (Vint (Int.neg y)). -Proof. - destruct x; intro y; simpl; auto; rewrite Int.sub_add_opp; auto. -Qed. - -Theorem sub_add_l: - forall v1 v2 i, sub (add v1 (Vint i)) v2 = add (sub v1 v2) (Vint i). -Proof. - destruct v1; destruct v2; intros; simpl; auto. - rewrite Int.sub_add_l. auto. - rewrite Int.sub_add_l. auto. - case (zeq b b0); intro. rewrite Int.sub_add_l. auto. reflexivity. -Qed. - -Theorem sub_add_r: - forall v1 v2 i, sub v1 (add v2 (Vint i)) = add (sub v1 v2) (Vint (Int.neg i)). -Proof. - destruct v1; destruct v2; intros; simpl; auto. - rewrite Int.sub_add_r. auto. - repeat rewrite Int.sub_add_opp. decEq. - repeat rewrite Int.add_assoc. decEq. apply Int.add_commut. - decEq. repeat rewrite Int.sub_add_opp. - rewrite Int.add_assoc. decEq. apply Int.neg_add_distr. - case (zeq b b0); intro. simpl. decEq. - repeat rewrite Int.sub_add_opp. rewrite Int.add_assoc. decEq. - apply Int.neg_add_distr. - reflexivity. -Qed. - -Theorem mul_commut: forall x y, mul x y = mul y x. -Proof. - destruct x; destruct y; simpl; auto. decEq. apply Int.mul_commut. -Qed. - -Theorem mul_assoc: forall x y z, mul (mul x y) z = mul x (mul y z). -Proof. - destruct x; destruct y; destruct z; simpl; auto. - decEq. apply Int.mul_assoc. -Qed. - -Theorem mul_add_distr_l: - forall x y z, mul (add x y) z = add (mul x z) (mul y z). -Proof. - destruct x; destruct y; destruct z; simpl; auto. - decEq. apply Int.mul_add_distr_l. -Qed. - - -Theorem mul_add_distr_r: - forall x y z, mul x (add y z) = add (mul x y) (mul x z). -Proof. - destruct x; destruct y; destruct z; simpl; auto. - decEq. apply Int.mul_add_distr_r. -Qed. - -Theorem mul_pow2: - forall x n logn, - Int.is_power2 n = Some logn -> - mul x (Vint n) = shl x (Vint logn). -Proof. - intros; destruct x; simpl; auto. - change 32 with (Z_of_nat wordsize). - rewrite (Int.is_power2_range _ _ H). decEq. apply Int.mul_pow2. auto. -Qed. - -Theorem mods_divs: - forall x y, mods x y = sub x (mul (divs x y) y). -Proof. - destruct x; destruct y; simpl; auto. - case (Int.eq i0 Int.zero); simpl. auto. decEq. apply Int.mods_divs. -Qed. - -Theorem modu_divu: - forall x y, modu x y = sub x (mul (divu x y) y). -Proof. - destruct x; destruct y; simpl; auto. - generalize (Int.eq_spec i0 Int.zero); - case (Int.eq i0 Int.zero); simpl. auto. - intro. decEq. apply Int.modu_divu. auto. -Qed. - -Theorem divs_pow2: - forall x n logn, - Int.is_power2 n = Some logn -> - divs x (Vint n) = shrx x (Vint logn). -Proof. - intros; destruct x; simpl; auto. - change 32 with (Z_of_nat wordsize). - rewrite (Int.is_power2_range _ _ H). - generalize (Int.eq_spec n Int.zero); - case (Int.eq n Int.zero); intro. - subst n. compute in H. discriminate. - decEq. apply Int.divs_pow2. auto. -Qed. - -Theorem divu_pow2: - forall x n logn, - Int.is_power2 n = Some logn -> - divu x (Vint n) = shru x (Vint logn). -Proof. - intros; destruct x; simpl; auto. - change 32 with (Z_of_nat wordsize). - rewrite (Int.is_power2_range _ _ H). - generalize (Int.eq_spec n Int.zero); - case (Int.eq n Int.zero); intro. - subst n. compute in H. discriminate. - decEq. apply Int.divu_pow2. auto. -Qed. - -Theorem modu_pow2: - forall x n logn, - Int.is_power2 n = Some logn -> - modu x (Vint n) = and x (Vint (Int.sub n Int.one)). -Proof. - intros; destruct x; simpl; auto. - generalize (Int.eq_spec n Int.zero); - case (Int.eq n Int.zero); intro. - subst n. compute in H. discriminate. - decEq. eapply Int.modu_and; eauto. -Qed. - -Theorem and_commut: forall x y, and x y = and y x. -Proof. - destruct x; destruct y; simpl; auto. decEq. apply Int.and_commut. -Qed. - -Theorem and_assoc: forall x y z, and (and x y) z = and x (and y z). -Proof. - destruct x; destruct y; destruct z; simpl; auto. - decEq. apply Int.and_assoc. -Qed. - -Theorem or_commut: forall x y, or x y = or y x. -Proof. - destruct x; destruct y; simpl; auto. decEq. apply Int.or_commut. -Qed. - -Theorem or_assoc: forall x y z, or (or x y) z = or x (or y z). -Proof. - destruct x; destruct y; destruct z; simpl; auto. - decEq. apply Int.or_assoc. -Qed. - -Theorem xor_commut: forall x y, xor x y = xor y x. -Proof. - destruct x; destruct y; simpl; auto. decEq. apply Int.xor_commut. -Qed. - -Theorem xor_assoc: forall x y z, xor (xor x y) z = xor x (xor y z). -Proof. - destruct x; destruct y; destruct z; simpl; auto. - decEq. apply Int.xor_assoc. -Qed. - -Theorem shl_mul: forall x y, Val.mul x (Val.shl Vone y) = Val.shl x y. -Proof. - destruct x; destruct y; simpl; auto. - case (Int.ltu i0 (Int.repr 32)); auto. - decEq. symmetry. apply Int.shl_mul. -Qed. - -Theorem shl_rolm: - forall x n, - Int.ltu n (Int.repr 32) = true -> - shl x (Vint n) = rolm x n (Int.shl Int.mone n). -Proof. - intros; destruct x; simpl; auto. - rewrite H. decEq. apply Int.shl_rolm. exact H. -Qed. - -Theorem shru_rolm: - forall x n, - Int.ltu n (Int.repr 32) = true -> - shru x (Vint n) = rolm x (Int.sub (Int.repr 32) n) (Int.shru Int.mone n). -Proof. - intros; destruct x; simpl; auto. - rewrite H. decEq. apply Int.shru_rolm. exact H. -Qed. - -Theorem shrx_carry: - forall x y, - add (shr x y) (shr_carry x y) = shrx x y. -Proof. - destruct x; destruct y; simpl; auto. - case (Int.ltu i0 (Int.repr 32)); auto. - simpl. decEq. apply Int.shrx_carry. -Qed. - -Theorem or_rolm: - forall x n m1 m2, - or (rolm x n m1) (rolm x n m2) = rolm x n (Int.or m1 m2). -Proof. - intros; destruct x; simpl; auto. - decEq. apply Int.or_rolm. -Qed. - -Theorem rolm_rolm: - forall x n1 m1 n2 m2, - rolm (rolm x n1 m1) n2 m2 = - rolm x (Int.and (Int.add n1 n2) (Int.repr 31)) - (Int.and (Int.rol m1 n2) m2). -Proof. - intros; destruct x; simpl; auto. - decEq. - replace (Int.and (Int.add n1 n2) (Int.repr 31)) - with (Int.modu (Int.add n1 n2) (Int.repr 32)). - apply Int.rolm_rolm. - change (Int.repr 31) with (Int.sub (Int.repr 32) Int.one). - apply Int.modu_and with (Int.repr 5). reflexivity. -Qed. - -Theorem rolm_zero: - forall x m, - rolm x Int.zero m = and x (Vint m). -Proof. - intros; destruct x; simpl; auto. decEq. apply Int.rolm_zero. -Qed. - -Theorem addf_commut: forall x y, addf x y = addf y x. -Proof. - destruct x; destruct y; simpl; auto. decEq. apply Float.addf_commut. -Qed. - -Lemma negate_cmp_mismatch: - forall c, - cmp_mismatch (negate_comparison c) = notbool(cmp_mismatch c). -Proof. - destruct c; reflexivity. -Qed. - -Theorem negate_cmp: - forall c x y, - cmp (negate_comparison c) x y = notbool (cmp c x y). -Proof. - destruct x; destruct y; simpl; auto. - rewrite Int.negate_cmp. apply notbool_negb_1. - case (Int.eq i Int.zero). apply negate_cmp_mismatch. reflexivity. - case (Int.eq i0 Int.zero). apply negate_cmp_mismatch. reflexivity. - case (zeq b b0); intro. - rewrite Int.negate_cmp. apply notbool_negb_1. - apply negate_cmp_mismatch. -Qed. - -Theorem negate_cmpu: - forall c x y, - cmpu (negate_comparison c) x y = notbool (cmpu c x y). -Proof. - destruct x; destruct y; simpl; auto. - rewrite Int.negate_cmpu. apply notbool_negb_1. - case (Int.eq i Int.zero). apply negate_cmp_mismatch. reflexivity. - case (Int.eq i0 Int.zero). apply negate_cmp_mismatch. reflexivity. - case (zeq b b0); intro. - rewrite Int.negate_cmpu. apply notbool_negb_1. - apply negate_cmp_mismatch. -Qed. - -Lemma swap_cmp_mismatch: - forall c, cmp_mismatch (swap_comparison c) = cmp_mismatch c. -Proof. - destruct c; reflexivity. -Qed. - -Theorem swap_cmp: - forall c x y, - cmp (swap_comparison c) x y = cmp c y x. -Proof. - destruct x; destruct y; simpl; auto. - rewrite Int.swap_cmp. auto. - case (Int.eq i Int.zero). apply swap_cmp_mismatch. auto. - case (Int.eq i0 Int.zero). apply swap_cmp_mismatch. auto. - case (zeq b b0); intro. - subst b0. rewrite zeq_true. rewrite Int.swap_cmp. auto. - rewrite zeq_false. apply swap_cmp_mismatch. auto. -Qed. - -Theorem swap_cmpu: - forall c x y, - cmpu (swap_comparison c) x y = cmpu c y x. -Proof. - destruct x; destruct y; simpl; auto. - rewrite Int.swap_cmpu. auto. - case (Int.eq i Int.zero). apply swap_cmp_mismatch. auto. - case (Int.eq i0 Int.zero). apply swap_cmp_mismatch. auto. - case (zeq b b0); intro. - subst b0. rewrite zeq_true. rewrite Int.swap_cmpu. auto. - rewrite zeq_false. apply swap_cmp_mismatch. auto. -Qed. - -Theorem negate_cmpf_eq: - forall v1 v2, notbool (cmpf Cne v1 v2) = cmpf Ceq v1 v2. -Proof. - destruct v1; destruct v2; simpl; auto. - rewrite Float.cmp_ne_eq. rewrite notbool_negb_1. - apply notbool_idem2. -Qed. - -Lemma or_of_bool: - forall b1 b2, or (of_bool b1) (of_bool b2) = of_bool (b1 || b2). -Proof. - destruct b1; destruct b2; reflexivity. -Qed. - -Theorem cmpf_le: - forall v1 v2, cmpf Cle v1 v2 = or (cmpf Clt v1 v2) (cmpf Ceq v1 v2). -Proof. - destruct v1; destruct v2; simpl; auto. - rewrite or_of_bool. decEq. apply Float.cmp_le_lt_eq. -Qed. - -Theorem cmpf_ge: - forall v1 v2, cmpf Cge v1 v2 = or (cmpf Cgt v1 v2) (cmpf Ceq v1 v2). -Proof. - destruct v1; destruct v2; simpl; auto. - rewrite or_of_bool. decEq. apply Float.cmp_ge_gt_eq. -Qed. - -Definition is_bool (v: val) := - v = Vundef \/ v = Vtrue \/ v = Vfalse. - -Lemma of_bool_is_bool: - forall b, is_bool (of_bool b). -Proof. - destruct b; unfold is_bool; simpl; tauto. -Qed. - -Lemma undef_is_bool: is_bool Vundef. -Proof. - unfold is_bool; tauto. -Qed. - -Lemma cmp_mismatch_is_bool: - forall c, is_bool (cmp_mismatch c). -Proof. - destruct c; simpl; unfold is_bool; tauto. -Qed. - -Lemma cmp_is_bool: - forall c v1 v2, is_bool (cmp c v1 v2). -Proof. - destruct v1; destruct v2; simpl; try apply undef_is_bool. - apply of_bool_is_bool. - case (Int.eq i Int.zero). apply cmp_mismatch_is_bool. apply undef_is_bool. - case (Int.eq i0 Int.zero). apply cmp_mismatch_is_bool. apply undef_is_bool. - case (zeq b b0); intro. apply of_bool_is_bool. apply cmp_mismatch_is_bool. -Qed. - -Lemma cmpu_is_bool: - forall c v1 v2, is_bool (cmpu c v1 v2). -Proof. - destruct v1; destruct v2; simpl; try apply undef_is_bool. - apply of_bool_is_bool. - case (Int.eq i Int.zero). apply cmp_mismatch_is_bool. apply undef_is_bool. - case (Int.eq i0 Int.zero). apply cmp_mismatch_is_bool. apply undef_is_bool. - case (zeq b b0); intro. apply of_bool_is_bool. apply cmp_mismatch_is_bool. -Qed. - -Lemma cmpf_is_bool: - forall c v1 v2, is_bool (cmpf c v1 v2). -Proof. - destruct v1; destruct v2; simpl; - apply undef_is_bool || apply of_bool_is_bool. -Qed. - -Lemma notbool_is_bool: - forall v, is_bool (notbool v). -Proof. - destruct v; simpl. - apply undef_is_bool. apply of_bool_is_bool. - apply undef_is_bool. unfold is_bool; tauto. -Qed. - -Lemma notbool_xor: - forall v, is_bool v -> v = xor (notbool v) Vone. -Proof. - intros. elim H; intro. - subst v. reflexivity. - elim H0; intro; subst v; reflexivity. -Qed. - -End Val. -- cgit