diff --git a/Makefile b/Makefile index ec9028ac01..cf444c32cc 100644 --- a/Makefile +++ b/Makefile @@ -58,7 +58,8 @@ FLOCQ=\ VLIB=Axioms.v Coqlib.v Intv.v Maps.v Heaps.v Lattice.v Ordered.v \ Iteration.v Integers.v Archi.v Fappli_IEEE_extra.v Floats.v \ Parmov.v UnionFind.v Wfsimpl.v \ - Postorder.v FSetAVLplus.v IntvSets.v Decidableplus.v BoolEqual.v + Postorder.v FSetAVLplus.v IntvSets.v Decidableplus.v BoolEqual.v \ + Symbols.v # Parts common to the front-ends and the back-end (in common/) diff --git a/Makefile.extr b/Makefile.extr index a1c2ef7c0a..e0697efc4f 100644 --- a/Makefile.extr +++ b/Makefile.extr @@ -50,6 +50,7 @@ INCLUDES=$(patsubst %,-I %, $(DIRS)) # Control of warnings: WARNINGS=-w +a-4-9-27 -strict-sequence -safe-string -warn-error +a #Deprication returns with ocaml 4.03 +extraction/%.cmi: WARNINGS +=-w -20-27-32..34-39-41-44..45 extraction/%.cmx: WARNINGS +=-w -20-27-32..34-39-41-44..45 extraction/%.cmo: WARNINGS +=-w -20-27-32..34-39-41-44..45 cparser/pre_parser.cmx: WARNINGS += -w -41 diff --git a/arm/AsmToJSON.ml b/arm/AsmToJSON.ml index 49cca9298e..a6084b8a8e 100644 --- a/arm/AsmToJSON.ml +++ b/arm/AsmToJSON.ml @@ -37,7 +37,7 @@ let mnemonic_names = [ "Padc"; "Padd"; "Padds"; "Pand";"Pannot"; "Pasr"; "Pb"; " type instruction_arg = | ALabel of positive - | Atom of positive + | Atom of ident | Data32 of Integers.Int.int | Data64 of Integers.Int64.int | DFreg of freg diff --git a/backend/Asmexpandaux.mli b/backend/Asmexpandaux.mli index 797eb10c6d..b903453fae 100644 --- a/backend/Asmexpandaux.mli +++ b/backend/Asmexpandaux.mli @@ -13,7 +13,6 @@ open Asm open AST -open BinNums (** Auxiliary functions for builtin expansion *) @@ -31,6 +30,6 @@ val set_current_function: coq_function -> unit (* Set the current function *) val get_current_function: unit -> coq_function (* Get the current function *) -val expand_debug: positive -> int -> (preg -> int) -> (instruction -> unit) -> instruction list -> unit +val expand_debug: ident -> int -> (preg -> int) -> (instruction -> unit) -> instruction list -> unit (* Expand builtin debug function. Takes the function id, the register number of the stackpointer, a function to get the dwarf mapping of varibale names and for the expansion of simple instructions *) diff --git a/backend/Cminor.v b/backend/Cminor.v index 11941da31a..c1a29fac47 100644 --- a/backend/Cminor.v +++ b/backend/Cminor.v @@ -185,7 +185,7 @@ Definition funsig (fd: fundef) := *) Definition genv := Genv.t fundef unit. -Definition env := PTree.t val. +Definition env := ATree.t val. (** The following functions build the initial local environment at function entry, binding parameters to the provided arguments and @@ -193,21 +193,21 @@ Definition env := PTree.t val. Fixpoint set_params (vl: list val) (il: list ident) {struct il} : env := match il, vl with - | i1 :: is, v1 :: vs => PTree.set i1 v1 (set_params vs is) - | i1 :: is, nil => PTree.set i1 Vundef (set_params nil is) - | _, _ => PTree.empty val + | i1 :: is, v1 :: vs => ATree.set i1 v1 (set_params vs is) + | i1 :: is, nil => ATree.set i1 Vundef (set_params nil is) + | _, _ => ATree.empty val end. Fixpoint set_locals (il: list ident) (e: env) {struct il} : env := match il with | nil => e - | i1 :: is => PTree.set i1 Vundef (set_locals is e) + | i1 :: is => ATree.set i1 Vundef (set_locals is e) end. Definition set_optvar (optid: option ident) (v: val) (e: env) : env := match optid with | None => e - | Some id => PTree.set id v e + | Some id => ATree.set id v e end. (** Continuations *) @@ -359,7 +359,7 @@ Variable m: mem. Inductive eval_expr: expr -> val -> Prop := | eval_Evar: forall id v, - PTree.get id e = Some v -> + ATree.get id e = Some v -> eval_expr (Evar id) v | eval_Econst: forall cst v, eval_constant sp cst = Some v -> @@ -447,7 +447,7 @@ Inductive step: state -> trace -> state -> Prop := | step_assign: forall f id a k sp e m v, eval_expr sp e m a v -> step (State f (Sassign id a) k sp e m) - E0 (State f Sskip k sp (PTree.set id v e) m) + E0 (State f Sskip k sp (ATree.set id v e) m) | step_store: forall f chunk addr a k sp e m vaddr v m', eval_expr sp e m addr vaddr -> @@ -673,7 +673,7 @@ with exec_stmt: | exec_Sassign: forall f sp e m id a v, eval_expr ge sp e m a v -> - exec_stmt f sp e m (Sassign id a) E0 (PTree.set id v e) m Out_normal + exec_stmt f sp e m (Sassign id a) E0 (ATree.set id v e) m Out_normal | exec_Sstore: forall f sp e m chunk addr a vaddr v m', eval_expr ge sp e m addr vaddr -> @@ -948,7 +948,7 @@ Proof. constructor. (* assign *) - exists (State f Sskip k sp (PTree.set id v e) m); split. + exists (State f Sskip k sp (ATree.set id v e) m); split. apply star_one. constructor. auto. constructor. diff --git a/backend/CminorSel.v b/backend/CminorSel.v index 96cb8ae662..bab806a228 100644 --- a/backend/CminorSel.v +++ b/backend/CminorSel.v @@ -165,7 +165,7 @@ Variable m: mem. Inductive eval_expr: letenv -> expr -> val -> Prop := | eval_Evar: forall le id v, - PTree.get id e = Some v -> + ATree.get id e = Some v -> eval_expr le (Evar id) v | eval_Eop: forall le op al vl v, eval_exprlist le al vl -> @@ -284,7 +284,7 @@ End EVAL_EXPR. Definition set_builtin_res (res: builtin_res ident) (v: val) (e: env) : env := match res with - | BR id => PTree.set id v e + | BR id => ATree.set id v e | _ => e end. @@ -348,7 +348,7 @@ Inductive step: state -> trace -> state -> Prop := | step_assign: forall f id a k sp e m v, eval_expr sp e m nil a v -> step (State f (Sassign id a) k sp e m) - E0 (State f Sskip k sp (PTree.set id v e) m) + E0 (State f Sskip k sp (ATree.set id v e) m) | step_store: forall f chunk addr al b k sp e m vl v vaddr m', eval_exprlist sp e m nil al vl -> diff --git a/backend/Debugvar.v b/backend/Debugvar.v index 1f36103017..4a4da6ea65 100644 --- a/backend/Debugvar.v +++ b/backend/Debugvar.v @@ -74,7 +74,7 @@ Fixpoint set_state (v: ident) (i: debuginfo) (s: avail) : avail := match s with | nil => (v, i) :: nil | (v', i') as vi' :: s' => - match Pos.compare v v' with + match Ident.compare v v' with | Eq => (v, i) :: s' | Lt => (v, i) :: s | Gt => vi' :: set_state v i s' @@ -85,7 +85,7 @@ Fixpoint remove_state (v: ident) (s: avail) : avail := match s with | nil => nil | (v', i') as vi' :: s' => - match Pos.compare v v' with + match Ident.compare v v' with | Eq => s' | Lt => s | Gt => vi' :: remove_state v s' @@ -157,7 +157,7 @@ Fixpoint join (s1: avail) (s2: avail) {struct s1} : avail := match s2 with | nil => nil | (v2, i2) as vi2 :: s2' => - match Pos.compare v1 v2 with + match Ident.compare v1 v2 with | Eq => if eq_debuginfo i1 i2 then vi1 :: join s1' s2' else join s1' s2' | Lt => join s1' s2 | Gt => join2 s2' @@ -296,7 +296,7 @@ Fixpoint diff (s1 s2: avail) {struct s1} : avail := match s2 with | nil => s1 | (v2, i2) :: s2' => - match Pos.compare v1 v2 with + match Ident.compare v1 v2 with | Eq => if eq_debuginfo i1 i2 then diff s1' s2' else vi1 :: diff s1' s2' | Lt => vi1 :: diff s1' s2 | Gt => diff2 s2' diff --git a/backend/Debugvarproof.v b/backend/Debugvarproof.v index d31c63ec74..cb4b69cf87 100644 --- a/backend/Debugvarproof.v +++ b/backend/Debugvarproof.v @@ -41,7 +41,8 @@ Remark diff_same: Proof. induction s as [ | [v i] s]; simpl. auto. - rewrite Pos.compare_refl. rewrite dec_eq_true. auto. + destruct (Ident.compare_spec v v); try solve [eelim Ident.lt_not_eq; eauto]. + rewrite dec_eq_true. auto. Qed. Remark delta_state_same: @@ -117,7 +118,7 @@ Qed. but establish some confidence in the availability analysis. *) Definition avail_above (v: ident) (s: avail) : Prop := - forall v' i', In (v', i') s -> Plt v v'. + forall v' i', In (v', i') s -> Ident.lt v v'. Inductive wf_avail: avail -> Prop := | wf_avail_nil: @@ -132,7 +133,7 @@ Lemma set_state_1: Proof. induction s as [ | [v' i'] s]; simpl. - auto. -- destruct (Pos.compare v v'); simpl; auto. +- destruct (Ident.compare v v'); simpl; auto. Qed. Lemma set_state_2: @@ -141,7 +142,7 @@ Lemma set_state_2: Proof. induction s as [ | [v1 i1] s]; simpl; intros. - contradiction. -- destruct (Pos.compare_spec v v1); simpl. +- destruct (Ident.compare_spec v v1); simpl. + subst v1. destruct H0. congruence. auto. + auto. + destruct H0; auto. @@ -155,14 +156,14 @@ Lemma set_state_3: Proof. induction 1; simpl; intros. - intuition congruence. -- destruct (Pos.compare_spec v v0); simpl in H1. +- destruct (Ident.compare_spec v v0); simpl in H1. + subst v0. destruct H1. inv H1; auto. right; split. - apply not_eq_sym. apply Plt_ne. eapply H; eauto. + apply not_eq_sym. apply Ident.lt_not_eq. eapply H; eauto. auto. + destruct H1. inv H1; auto. - destruct H1. inv H1. right; split; auto. apply not_eq_sym. apply Plt_ne. auto. - right; split; auto. apply not_eq_sym. apply Plt_ne. apply Plt_trans with v0; eauto. -+ destruct H1. inv H1. right; split; auto. apply Plt_ne. auto. + destruct H1. inv H1. right; split; auto. apply not_eq_sym. apply Ident.lt_not_eq. auto. + right; split; auto. apply not_eq_sym. apply Ident.lt_not_eq. apply Ident.lt_trans with v0; eauto. ++ destruct H1. inv H1. right; split; auto. apply Ident.lt_not_eq. auto. destruct IHwf_avail as [A | [A B]]; auto. Qed. @@ -171,11 +172,11 @@ Lemma wf_set_state: Proof. induction 1; simpl. - constructor. red; simpl; tauto. constructor. -- destruct (Pos.compare_spec v v0). +- destruct (Ident.compare_spec v v0). + subst v0. constructor; auto. + constructor. red; simpl; intros. destruct H2. - inv H2. auto. apply Plt_trans with v0; eauto. + inv H2. auto. apply Ident.lt_trans with v0; eauto. constructor; auto. + constructor. red; intros. exploit set_state_3. eexact H0. eauto. intros [[A B] | [A B]]; subst; eauto. @@ -187,11 +188,11 @@ Lemma remove_state_1: Proof. induction 1; simpl; red; intros. - auto. -- destruct (Pos.compare_spec v v0); simpl in *. -+ subst v0. elim (Plt_strict v); eauto. -+ destruct H1. inv H1. elim (Plt_strict v); eauto. - elim (Plt_strict v). apply Plt_trans with v0; eauto. -+ destruct H1. inv H1. elim (Plt_strict v); eauto. tauto. +- destruct (Ident.compare_spec v v0); simpl in *. ++ subst v0. elim (Ident.lt_not_eq v v); eauto. ++ destruct H1. inv H1. elim (Ident.lt_not_eq v v); eauto. + elim (Ident.lt_not_eq v v); eauto. apply Ident.lt_trans with v0; eauto. ++ destruct H1. inv H1. elim (Ident.lt_not_eq v v); eauto. tauto. Qed. Lemma remove_state_2: @@ -199,7 +200,7 @@ Lemma remove_state_2: Proof. induction s as [ | [v1 i1] s]; simpl; intros. - auto. -- destruct (Pos.compare_spec v v1); simpl. +- destruct (Ident.compare_spec v v1); simpl. + subst v1. destruct H0. congruence. auto. + auto. + destruct H0; auto. @@ -210,11 +211,11 @@ Lemma remove_state_3: Proof. induction 1; simpl; intros. - contradiction. -- destruct (Pos.compare_spec v v0); simpl in H1. -+ subst v0. split; auto. apply not_eq_sym; apply Plt_ne; eauto. -+ destruct H1. inv H1. split; auto. apply not_eq_sym; apply Plt_ne; eauto. - split; auto. apply not_eq_sym; apply Plt_ne. apply Plt_trans with v0; eauto. -+ destruct H1. inv H1. split; auto. apply Plt_ne; auto. +- destruct (Ident.compare_spec v v0); simpl in H1. ++ subst v0. split; auto. apply not_eq_sym; apply Ident.lt_not_eq; eauto. ++ destruct H1. inv H1. split; auto. apply not_eq_sym; apply Ident.lt_not_eq; eauto. + split; auto. apply not_eq_sym; apply Ident.lt_not_eq. apply Ident.lt_trans with v0; eauto. ++ destruct H1. inv H1. split; auto. apply Ident.lt_not_eq; auto. destruct IHwf_avail as [A B] ; auto. Qed. @@ -223,7 +224,7 @@ Lemma wf_remove_state: Proof. induction 1; simpl. - constructor. -- destruct (Pos.compare_spec v v0). +- destruct (Ident.compare_spec v v0). + auto. + constructor; auto. + constructor; auto. red; intros. @@ -246,12 +247,22 @@ Lemma join_1: Proof. induction 1; simpl; try tauto; induction 1; simpl; intros I1 I2; auto. destruct I1, I2. -- inv H3; inv H4. rewrite Pos.compare_refl. rewrite dec_eq_true; auto with coqlib. +- inv H3; inv H4. + destruct (Ident.compare_spec v v); try solve [eelim Ident.lt_not_eq; eauto]. + rewrite dec_eq_true; auto with coqlib. - inv H3. - assert (L: Plt v1 v) by eauto. apply Pos.compare_gt_iff in L. rewrite L. auto. + assert (L: Ident.lt v1 v) by eauto. + destruct (Ident.compare_spec v v1). + + eelim Ident.lt_not_eq; eauto. + + elim (Ident.lt_not_eq v v); eauto using Ident.lt_trans. + + auto. - inv H4. - assert (L: Plt v0 v) by eauto. apply Pos.compare_lt_iff in L. rewrite L. apply IHwf_avail. constructor; auto. auto. auto with coqlib. -- destruct (Pos.compare v0 v1). + assert (L: Ident.lt v0 v) by eauto. + destruct (Ident.compare_spec v0 v). + + eelim Ident.lt_not_eq; eauto. + + apply IHwf_avail. constructor; auto. auto. auto with coqlib. + + elim (Ident.lt_not_eq v v); eauto using Ident.lt_trans. +- destruct (Ident.compare v0 v1). + destruct (eq_debuginfo i0 i1); auto with coqlib. + apply IHwf_avail; auto with coqlib. constructor; auto. + eauto. @@ -262,7 +273,7 @@ Lemma join_2: In (v, i) (join s1 s2) -> In (v, i) s1 /\ In (v, i) s2. Proof. induction 1; simpl; try tauto; induction 1; simpl; intros I; try tauto. - destruct (Pos.compare_spec v0 v1). + destruct (Ident.compare_spec v0 v1). - subst v1. destruct (eq_debuginfo i0 i1). + subst i1. destruct I. auto. exploit IHwf_avail; eauto. tauto. + exploit IHwf_avail; eauto. tauto. @@ -275,7 +286,7 @@ Lemma wf_join: forall s1, wf_avail s1 -> forall s2, wf_avail s2 -> wf_avail (join s1 s2). Proof. induction 1; simpl; induction 1; simpl; try constructor. - destruct (Pos.compare_spec v v0). + destruct (Ident.compare_spec v v0). - subst v0. destruct (eq_debuginfo i i0); auto. constructor; auto. red; intros. apply join_2 in H3; auto. destruct H3. eauto. - apply IHwf_avail. constructor; auto. diff --git a/backend/Inlining.v b/backend/Inlining.v index f7ee416640..faf61d5713 100644 --- a/backend/Inlining.v +++ b/backend/Inlining.v @@ -23,9 +23,9 @@ Require Import Op Registers RTL. inlining, as determined by the external heuristic [should_inline]. *) -Definition funenv : Type := PTree.t function. +Definition funenv : Type := ATree.t function. -Definition size_fenv (fenv: funenv) := PTree_Properties.cardinal fenv. +Definition size_fenv (fenv: funenv) := ATree_Properties.cardinal fenv. Parameter inlining_info: Type. (* abstract type, implemented on the Caml side *) @@ -37,15 +37,15 @@ Definition add_globdef (io: inlining_info) (fenv: funenv) (idg: ident * globdef match idg with | (id, Gfun (Internal f)) => if should_inline io id f - then PTree.set id f fenv - else PTree.remove id fenv + then ATree.set id f fenv + else ATree.remove id fenv | (id, _) => - PTree.remove id fenv + ATree.remove id fenv end. Definition funenv_program (p: program) : funenv := let io := inlining_analysis p in - List.fold_left (add_globdef io) p.(prog_defs) (PTree.empty function). + List.fold_left (add_globdef io) p.(prog_defs) (ATree.empty function). (** State monad *) @@ -291,12 +291,12 @@ Variable rec: forall fenv', (size_fenv fenv' < size_fenv fenv)%nat -> context -> Inductive inline_decision (ros: reg + ident) : Type := | Cannot_inline - | Can_inline (id: ident) (f: function) (P: ros = inr reg id) (Q: fenv!id = Some f). + | Can_inline (id: ident) (f: function) (P: ros = inr reg id) (Q: fenv$id = Some f). Program Definition can_inline (ros: reg + ident): inline_decision ros := match ros with | inl r => Cannot_inline _ - | inr id => match fenv!id with Some f => Can_inline _ id f _ _ | None => Cannot_inline _ end + | inr id => match fenv$id with Some f => Can_inline _ id f _ _ | None => Cannot_inline _ end end. (** Inlining of a call to function [f]. An appropriate context is @@ -304,7 +304,7 @@ Program Definition can_inline (ros: reg + ident): inline_decision ros := are inserted to copy the arguments of the call to the parameters of [f]. *) Definition inline_function (ctx: context) (id: ident) (f: function) - (P: PTree.get id fenv = Some f) + (P: ATree.get id fenv = Some f) (args: list reg) (retpc: node) (retreg: reg) : mon node := let npc := max_pc_function f in @@ -312,21 +312,21 @@ Definition inline_function (ctx: context) (id: ident) (f: function) do dpc <- reserve_nodes npc; do dreg <- reserve_regs nreg; let ctx' := callcontext ctx dpc dreg nreg f.(fn_stacksize) retpc retreg in - do x <- rec (PTree.remove id fenv) (PTree_Properties.cardinal_remove P) ctx' f; + do x <- rec (ATree.remove id fenv) (ATree_Properties.cardinal_remove P) ctx' f; add_moves (sregs ctx args) (sregs ctx' f.(fn_params)) (spc ctx' f.(fn_entrypoint)). (** Inlining of a tail call to function [f]. Similar to [inline_function], but the new context is different. *) Definition inline_tail_function (ctx: context) (id: ident) (f: function) - (P: PTree.get id fenv = Some f) + (P: ATree.get id fenv = Some f) (args: list reg): mon node := let npc := max_pc_function f in let nreg := max_reg_function f in do dpc <- reserve_nodes npc; do dreg <- reserve_regs nreg; let ctx' := tailcontext ctx dpc dreg nreg f.(fn_stacksize) in - do x <- rec (PTree.remove id fenv) (PTree_Properties.cardinal_remove P) ctx' f; + do x <- rec (ATree.remove id fenv) (ATree_Properties.cardinal_remove P) ctx' f; add_moves (sregs ctx args) (sregs ctx' f.(fn_params)) (spc ctx' f.(fn_entrypoint)). (** The instruction generated for a [Ireturn] instruction found in an diff --git a/backend/Inliningaux.ml b/backend/Inliningaux.ml index 842e0c9310..c145d00821 100644 --- a/backend/Inliningaux.ml +++ b/backend/Inliningaux.ml @@ -15,38 +15,38 @@ open Datatypes open FSetAVL open Maps open Op -open Ordered +open Symbols open !RTL -module PSet = Make(OrderedPositive) +module ASet = Make(OrderedIdent) type inlining_info = { - call_cnt : int PTree.t; (* Count the number of direct calls to a function *) - addr_taken : PSet.t; (* The set of globals which have their address taken *) + call_cnt : int ATree.t; (* Count the number of direct calls to a function *) + addr_taken : ASet.t; (* The set of globals which have their address taken *) } let empty_inlining_info = { - call_cnt = PTree.empty; - addr_taken = PSet.empty; + call_cnt = ATree.empty; + addr_taken = ASet.empty; } let call_count id io = - match PTree.get id io.call_cnt with + match ATree.get id io.call_cnt with | Some cnt -> cnt | None -> 0 let called id io = - let call_cnt = PTree.set id (1 + call_count id io) io.call_cnt in + let call_cnt = ATree.set id (1 + call_count id io) io.call_cnt in { io with call_cnt = call_cnt } let address_taken id io = - PSet.mem id io.addr_taken + ASet.mem id io.addr_taken let rec used_id io ids = match ids with | [] -> io | id::ids -> - used_id {io with addr_taken = PSet.add id io.addr_taken} ids + used_id {io with addr_taken = ASet.add id io.addr_taken} ids let used_in_globvar io gv = let used_in_init_data io = function diff --git a/backend/Inliningproof.v b/backend/Inliningproof.v index 2dcb895687..3c2cf0e1d1 100644 --- a/backend/Inliningproof.v +++ b/backend/Inliningproof.v @@ -392,7 +392,7 @@ Lemma find_inlined_function: forall fenv id rs fd f, fenv_compat prog fenv -> find_function ge (inr id) rs = Some fd -> - fenv!id = Some f -> + fenv$id = Some f -> fd = Internal f. Proof. intros. diff --git a/backend/Inliningspec.v b/backend/Inliningspec.v index c345c94260..22015601e0 100644 --- a/backend/Inliningspec.v +++ b/backend/Inliningspec.v @@ -24,39 +24,39 @@ Require Import Inlining. Definition fenv_compat (p: program) (fenv: funenv) : Prop := forall id f, - fenv!id = Some f -> (prog_defmap p)!id = Some (Gfun (Internal f)). + fenv$id = Some f -> (prog_defmap p)$id = Some (Gfun (Internal f)). Lemma funenv_program_compat: forall p, fenv_compat p (funenv_program p). Proof. - set (P := fun (dm: PTree.t (globdef fundef unit)) (fenv: funenv) => + set (P := fun (dm: ATree.t (globdef fundef unit)) (fenv: funenv) => forall id f, - fenv!id = Some f -> dm!id = Some (Gfun (Internal f))). + fenv$id = Some f -> dm$id = Some (Gfun (Internal f))). assert (REMOVE: forall dm fenv id g, P dm fenv -> - P (PTree.set id g dm) (PTree.remove id fenv)). - { unfold P; intros. rewrite PTree.grspec in H0. destruct (PTree.elt_eq id0 id). + P (ATree.set id g dm) (ATree.remove id fenv)). + { unfold P; intros. rewrite ATree.grspec in H0. destruct (ATree.elt_eq id0 id). discriminate. - rewrite PTree.gso; auto. + rewrite ATree.gso; auto. } assert (ADD: forall io dm fenv idg, P dm fenv -> - P (PTree.set (fst idg) (snd idg) dm) (add_globdef io fenv idg)). + P (ATree.set (fst idg) (snd idg) dm) (add_globdef io fenv idg)). { intros io dm fenv [id g]; simpl; intros. destruct g as [ [f|ef] | v]; auto. destruct (should_inline io id f); auto. - red; intros. rewrite ! PTree.gsspec in *. - destruct (peq id0 id); auto. inv H0; auto. + red; intros. rewrite ! ATree.gsspec in *. + destruct (ATree.elt_eq id0 id); auto. inv H0; auto. } assert (REC: forall p l dm fenv, P dm fenv -> - P (fold_left (fun x idg => PTree.set (fst idg) (snd idg) x) l dm) + P (fold_left (fun x idg => ATree.set (fst idg) (snd idg) x) l dm) (fold_left (add_globdef p) l fenv)). { induction l; simpl; intros. - auto. - apply IHl. apply ADD; auto. } - intros. apply REC. red; intros. rewrite PTree.gempty in H; discriminate. + intros. apply REC. red; intros. rewrite ATree.gempty in H; discriminate. Qed. Lemma fenv_compat_linkorder: @@ -283,7 +283,7 @@ Inductive tr_instr: context -> node -> instruction -> code -> Prop := tr_instr ctx pc (Icall sg ros args res s) c | tr_call_inlined:forall ctx pc sg id args res s c f pc1 ctx', Ple res ctx.(mreg) -> - fenv!id = Some f -> + fenv$id = Some f -> c!(spc ctx pc) = Some(Inop pc1) -> tr_moves c pc1 (sregs ctx args) (sregs ctx' f.(fn_params)) (spc ctx' f.(fn_entrypoint)) -> tr_funbody ctx' f c -> @@ -300,7 +300,7 @@ Inductive tr_instr: context -> node -> instruction -> code -> Prop := ctx.(retinfo) = Some(s, res) -> tr_instr ctx pc (Itailcall sg ros args) c | tr_tailcall_inlined: forall ctx pc sg id args c f pc1 ctx', - fenv!id = Some f -> + fenv$id = Some f -> c!(spc ctx pc) = Some(Inop pc1) -> tr_moves c pc1 (sregs ctx args) (sregs ctx' f.(fn_params)) (spc ctx' f.(fn_entrypoint)) -> tr_funbody ctx' f c -> @@ -337,7 +337,7 @@ with tr_funbody: context -> function -> code -> Prop := tr_funbody ctx f c. Definition fenv_agree (fe: funenv) : Prop := - forall id f, fe!id = Some f -> fenv!id = Some f. + forall id f, fe$id = Some f -> fenv$id = Some f. Section EXPAND_INSTR. @@ -504,7 +504,7 @@ Proof. intros. rewrite S1. eapply set_instr_other; eauto. unfold node; xomega. xomega. xomega. eapply rec_spec; eauto. - red; intros. rewrite PTree.grspec in H. destruct (PTree.elt_eq id0 id); try discriminate. auto. + red; intros. rewrite ATree.grspec in H. destruct (ATree.elt_eq id0 id); try discriminate. auto. simpl. subst s2; simpl in *; xomega. simpl. subst s3; simpl in *; xomega. simpl. xomega. @@ -534,7 +534,7 @@ Proof. eapply add_moves_spec; eauto. intros. rewrite S1. eapply set_instr_other; eauto. unfold node; xomega. xomega. xomega. eapply rec_spec; eauto. - red; intros. rewrite PTree.grspec in H. destruct (PTree.elt_eq id0 id); try discriminate. auto. + red; intros. rewrite ATree.grspec in H. destruct (ATree.elt_eq id0 id); try discriminate. auto. simpl. subst s3; simpl in *. subst s2; simpl in *. xomega. simpl. subst s3; simpl in *; xomega. simpl. xomega. diff --git a/backend/NeedDomain.v b/backend/NeedDomain.v index d431f3d8c3..a507855add 100644 --- a/backend/NeedDomain.v +++ b/backend/NeedDomain.v @@ -1093,7 +1093,7 @@ Qed. Inductive nmem : Type := | NMemDead - | NMem (stk: ISet.t) (gl: PTree.t ISet.t). + | NMem (stk: ISet.t) (gl: ATree.t ISet.t). (** Interpretation of [nmem]: - [NMemDead]: all memory locations are unused (dead). Acts as bottom. @@ -1112,19 +1112,19 @@ Inductive nlive: nmem -> block -> Z -> Prop := (STK: b = sp -> ~ISet.In ofs stk) (GL: forall id iv, Genv.find_symbol ge id = Some b -> - gl!id = Some iv -> + gl$id = Some iv -> ~ISet.In ofs iv), nlive (NMem stk gl) b ofs. (** All locations are live *) -Definition nmem_all := NMem ISet.empty (PTree.empty _). +Definition nmem_all := NMem ISet.empty (ATree.empty _). Lemma nlive_all: forall b ofs, nlive nmem_all b ofs. Proof. intros; constructor; intros. apply ISet.In_empty. - rewrite PTree.gempty in H0; discriminate. + rewrite ATree.gempty in H0; discriminate. Qed. (** Add a range of live locations to [nm]. The range starts at @@ -1136,12 +1136,12 @@ Definition nmem_add (nm: nmem) (p: aptr) (sz: Z) : nmem := | NMem stk gl => match p with | Gl id ofs => - match gl!id with - | Some iv => NMem stk (PTree.set id (ISet.remove (Ptrofs.unsigned ofs) (Ptrofs.unsigned ofs + sz) iv) gl) + match gl$id with + | Some iv => NMem stk (ATree.set id (ISet.remove (Ptrofs.unsigned ofs) (Ptrofs.unsigned ofs + sz) iv) gl) | None => nm end | Glo id => - NMem stk (PTree.remove id gl) + NMem stk (ATree.remove id gl) | Stk ofs => NMem (ISet.remove (Ptrofs.unsigned ofs) (Ptrofs.unsigned ofs + sz) stk) gl | Stack => @@ -1162,11 +1162,11 @@ Proof. inv H1; try (apply nlive_all). - (* Gl id ofs *) assert (Genv.find_symbol ge id = Some b) by (eapply H; eauto). - destruct gl!id as [iv|] eqn:NG. + destruct gl$id as [iv|] eqn:NG. + constructor; simpl; intros. congruence. assert (id0 = id) by (eapply Genv.genv_vars_inj; eauto). subst id0. - rewrite PTree.gss in H5. inv H5. rewrite ISet.In_remove. + rewrite ATree.gss in H5. inv H5. rewrite ISet.In_remove. intros [A B]. elim A; auto. + constructor; simpl; intros. congruence. @@ -1177,7 +1177,7 @@ Proof. constructor; simpl; intros. congruence. assert (id0 = id) by (eapply Genv.genv_vars_inj; eauto). subst id0. - rewrite PTree.grs in H5. congruence. + rewrite ATree.grs in H5. congruence. - (* Stk ofs *) constructor; simpl; intros. rewrite ISet.In_remove. intros [A B]. elim A; auto. @@ -1194,14 +1194,14 @@ Lemma incl_nmem_add: Proof. intros. inversion H; subst. unfold nmem_add; destruct p; try (apply nlive_all). - (* Gl id ofs *) - destruct gl!id as [iv|] eqn:NG. + destruct gl$id as [iv|] eqn:NG. + split; simpl; intros. auto. - rewrite PTree.gsspec in H1. destruct (peq id0 id); eauto. inv H1. + rewrite ATree.gsspec in H1. destruct (ATree.elt_eq id0 id); eauto. inv H1. rewrite ISet.In_remove. intros [P Q]. eelim GL; eauto. + auto. - (* Glo id *) split; simpl; intros. auto. - rewrite PTree.grspec in H1. destruct (PTree.elt_eq id0 id). congruence. eauto. + rewrite ATree.grspec in H1. destruct (ATree.elt_eq id0 id). congruence. eauto. - (* Stk ofs *) split; simpl; intros. rewrite ISet.In_remove. intros [P Q]. eelim STK; eauto. @@ -1222,11 +1222,11 @@ Definition nmem_remove (nm: nmem) (p: aptr) (sz: Z) : nmem := match p with | Gl id ofs => let iv' := - match gl!id with + match gl$id with | Some iv => ISet.add (Ptrofs.unsigned ofs) (Ptrofs.unsigned ofs + sz) iv | None => ISet.interval (Ptrofs.unsigned ofs) (Ptrofs.unsigned ofs + sz) end in - NMem stk (PTree.set id iv' gl) + NMem stk (ATree.set id iv' gl) | Stk ofs => NMem (ISet.add (Ptrofs.unsigned ofs) (Ptrofs.unsigned ofs + sz) stk) gl | _ => nm @@ -1244,7 +1244,7 @@ Lemma nlive_remove: Proof. intros. inversion H2; subst. unfold nmem_remove; inv H1; auto. - (* Gl id ofs *) - set (iv' := match gl!id with + set (iv' := match gl$id with | Some iv => ISet.add (Ptrofs.unsigned ofs) (Ptrofs.unsigned ofs + sz) iv | None => @@ -1253,8 +1253,8 @@ Proof. end). assert (Genv.find_symbol ge id = Some b) by (eapply H; eauto). split; simpl; auto; intros. - rewrite PTree.gsspec in H6. destruct (peq id0 id). -+ inv H6. destruct H3. congruence. destruct gl!id as [iv0|] eqn:NG. + rewrite ATree.gsspec in H6. destruct (ATree.elt_eq id0 id). ++ inv H6. destruct H3. congruence. destruct gl$id as [iv0|] eqn:NG. unfold iv'; rewrite ISet.In_add. intros [P|P]. omega. eelim GL; eauto. unfold iv'; rewrite ISet.In_interval. omega. + eauto. @@ -1273,7 +1273,7 @@ Definition nmem_contains (nm: nmem) (p: aptr) (sz: Z) := | NMem stk gl => match p with | Gl id ofs => - match gl!id with + match gl$id with | Some iv => negb (ISet.contains (Ptrofs.unsigned ofs) (Ptrofs.unsigned ofs + sz) iv) | None => true end @@ -1296,7 +1296,7 @@ Proof. inv H1; try discriminate. - (* Gl id ofs *) assert (Genv.find_symbol ge id = Some b) by (eapply H; eauto). - destruct gl!id as [iv|] eqn:HG; inv H2. + destruct gl$id as [iv|] eqn:HG; inv H2. destruct (ISet.contains (Ptrofs.unsigned ofs) (Ptrofs.unsigned ofs + sz) iv) eqn:IC; try discriminate. rewrite ISet.contains_spec in IC. eelim GL; eauto. - (* Stk ofs *) @@ -1309,14 +1309,14 @@ Qed. a [Ireturn] or [Itailcall] instruction. *) Definition nmem_dead_stack (sz: Z) := - NMem (ISet.interval 0 sz) (PTree.empty _). + NMem (ISet.interval 0 sz) (ATree.empty _). Lemma nlive_dead_stack: forall sz b' i, b' <> sp \/ ~(0 <= i < sz) -> nlive (nmem_dead_stack sz) b' i. Proof. intros; constructor; simpl; intros. - rewrite ISet.In_interval. intuition. -- rewrite PTree.gempty in H1; discriminate. +- rewrite ATree.gempty in H1; discriminate. Qed. (** Least upper bound *) @@ -1327,7 +1327,7 @@ Definition nmem_lub (nm1 nm2: nmem) : nmem := | _, NMemDead => nm1 | NMem stk1 gl1, NMem stk2 gl2 => NMem (ISet.inter stk1 stk2) - (PTree.combine + (ATree.combine (fun o1 o2 => match o1, o2 with | Some iv1, Some iv2 => Some(ISet.inter iv1 iv2) @@ -1342,9 +1342,9 @@ Proof. intros. inversion H; subst. destruct nm2; simpl. auto. constructor; simpl; intros. - rewrite ISet.In_inter. intros [P Q]. eelim STK; eauto. -- rewrite PTree.gcombine in H1 by auto. - destruct gl!id as [iv1|] eqn:NG1; try discriminate; - destruct gl0!id as [iv2|] eqn:NG2; inv H1. +- rewrite ATree.gcombine in H1 by auto. + destruct gl$id as [iv1|] eqn:NG1; try discriminate; + destruct gl0$id as [iv2|] eqn:NG2; inv H1. rewrite ISet.In_inter. intros [P Q]. eelim GL; eauto. Qed. @@ -1354,9 +1354,9 @@ Proof. intros. inversion H; subst. destruct nm1; simpl. auto. constructor; simpl; intros. - rewrite ISet.In_inter. intros [P Q]. eelim STK; eauto. -- rewrite PTree.gcombine in H1 by auto. - destruct gl0!id as [iv1|] eqn:NG1; try discriminate; - destruct gl!id as [iv2|] eqn:NG2; inv H1. +- rewrite ATree.gcombine in H1 by auto. + destruct gl0$id as [iv1|] eqn:NG1; try discriminate; + destruct gl$id as [iv2|] eqn:NG2; inv H1. rewrite ISet.In_inter. intros [P Q]. eelim GL; eauto. Qed. @@ -1365,7 +1365,7 @@ Qed. Definition nmem_beq (nm1 nm2: nmem) : bool := match nm1, nm2 with | NMemDead, NMemDead => true - | NMem stk1 gl1, NMem stk2 gl2 => ISet.beq stk1 stk2 && PTree.beq ISet.beq gl1 gl2 + | NMem stk1 gl1, NMem stk2 gl2 => ISet.beq stk1 stk2 && ATree.beq ISet.beq gl1 gl2 | _, _ => false end. @@ -1377,13 +1377,13 @@ Proof. unfold nmem_beq; intros. destruct nm1 as [ | stk1 gl1]; destruct nm2 as [ | stk2 gl2]; try discriminate. - split; intros L; inv L. -- InvBooleans. rewrite ISet.beq_spec in H0. rewrite PTree.beq_correct in H1. +- InvBooleans. rewrite ISet.beq_spec in H0. rewrite ATree.beq_correct in H1. split; intros L; inv L; constructor; intros. + rewrite <- H0. eauto. -+ specialize (H1 id). rewrite H2 in H1. destruct gl1!id as [iv1|] eqn: NG; try contradiction. ++ specialize (H1 id). rewrite H2 in H1. destruct gl1$id as [iv1|] eqn: NG; try contradiction. rewrite ISet.beq_spec in H1. rewrite <- H1. eauto. + rewrite H0. eauto. -+ specialize (H1 id). rewrite H2 in H1. destruct gl2!id as [iv2|] eqn: NG; try contradiction. ++ specialize (H1 id). rewrite H2 in H1. destruct gl2$id as [iv2|] eqn: NG; try contradiction. rewrite ISet.beq_spec in H1. rewrite H1. eauto. Qed. diff --git a/backend/PrintAsmaux.ml b/backend/PrintAsmaux.ml index d985d5a41d..73b5462516 100644 --- a/backend/PrintAsmaux.ml +++ b/backend/PrintAsmaux.ml @@ -25,10 +25,10 @@ module type TARGET = val print_prologue: out_channel -> unit val print_epilogue: out_channel -> unit val print_align: out_channel -> int -> unit - val print_comm_symb: out_channel -> Z.t -> P.t -> int -> unit - val print_var_info: out_channel -> P.t -> unit - val print_fun_info: out_channel -> P.t -> unit - val get_section_names: P.t -> section_name * section_name * section_name + val print_comm_symb: out_channel -> Z.t -> ident -> int -> unit + val print_var_info: out_channel -> ident -> unit + val print_fun_info: out_channel -> ident -> unit + val get_section_names: ident -> section_name * section_name * section_name val print_file_line: out_channel -> string -> int -> unit val print_optional_fun_info: out_channel -> unit val cfi_startproc: out_channel -> unit @@ -39,7 +39,7 @@ module type TARGET = val section: out_channel -> section_name -> unit val name_of_section: section_name -> string val comment: string - val symbol: out_channel -> P.t -> unit + val symbol: out_channel -> ident -> unit val default_falignment: int val label: out_channel -> int -> unit val address: string diff --git a/backend/RTLgen.v b/backend/RTLgen.v index 9d7a850699..cdb5c0ddb4 100644 --- a/backend/RTLgen.v +++ b/backend/RTLgen.v @@ -36,7 +36,7 @@ Local Open Scope string_scope. during translation of expressions, when crossing a [Elet] binding. *) Record mapping: Type := mkmapping { - map_vars: PTree.t reg; + map_vars: ATree.t reg; map_letvars: list reg }. @@ -290,11 +290,11 @@ Definition new_reg : mon reg := (** ** Operations on mappings *) Definition init_mapping : mapping := - mkmapping (PTree.empty reg) nil. + mkmapping (ATree.empty reg) nil. Definition add_var (map: mapping) (name: ident) : mon (reg * mapping) := do r <- new_reg; - ret (r, mkmapping (PTree.set name r map.(map_vars)) + ret (r, mkmapping (ATree.set name r map.(map_vars)) map.(map_letvars)). Fixpoint add_vars (map: mapping) (names: list ident) @@ -308,7 +308,7 @@ Fixpoint add_vars (map: mapping) (names: list ident) end. Definition find_var (map: mapping) (name: ident) : mon reg := - match PTree.get name map.(map_vars) with + match ATree.get name map.(map_vars) with | None => error (Errors.MSG "RTLgen: unbound variable " :: Errors.CTX name :: nil) | Some r => ret r end. @@ -551,7 +551,7 @@ Parameter more_likely: condexpr -> stmt -> stmt -> bool. [nlist] if it terminates by an [exit n] construct. [rret] is the register where the return value of the function must be stored, if any. *) -Definition labelmap : Type := PTree.t node. +Definition labelmap : Type := ATree.t node. Fixpoint transl_stmt (map: mapping) (s: stmt) (nd: node) (nexits: list node) (ngoto: labelmap) (nret: node) (rret: option reg) @@ -628,7 +628,7 @@ Fixpoint transl_stmt (map: mapping) (s: stmt) (nd: node) end | Slabel lbl s' => do ns <- transl_stmt map s' nd nexits ngoto nret rret; - match ngoto!lbl with + match ngoto$lbl with | None => error (Errors.msg "RTLgen: unbound label") | Some n => do xx <- @@ -638,7 +638,7 @@ Fixpoint transl_stmt (map: mapping) (s: stmt) (nd: node) ret ns end | Sgoto lbl => - match ngoto!lbl with + match ngoto$lbl with | None => error (Errors.MSG "Undefined defined label " :: Errors.CTX lbl :: nil) | Some n => ret n @@ -650,7 +650,7 @@ Fixpoint transl_stmt (map: mapping) (s: stmt) (nd: node) Definition alloc_label (lbl: Cminor.label) (maps: labelmap * state) : labelmap * state := let (map, s) := maps in let n := s.(st_nextnode) in - (PTree.set lbl n map, + (ATree.set lbl n map, mkstate s.(st_nextreg) (Pos.succ s.(st_nextnode)) s.(st_code) (reserve_instr_wf s)). Fixpoint reserve_labels (s: stmt) (ms: labelmap * state) @@ -682,7 +682,7 @@ Definition transl_fun (f: CminorSel.function) (ngoto: labelmap): mon (node * lis ret (nentry, rparams). Definition transl_function (f: CminorSel.function) : Errors.res RTL.function := - let (ngoto, s0) := reserve_labels f.(fn_body) (PTree.empty node, init_state) in + let (ngoto, s0) := reserve_labels f.(fn_body) (ATree.empty node, init_state) in match transl_fun f ngoto s0 with | Error msg => Errors.Error msg | OK (nentry, rparams) s i => diff --git a/backend/RTLgenproof.v b/backend/RTLgenproof.v index 072db1387a..f507accace 100644 --- a/backend/RTLgenproof.v +++ b/backend/RTLgenproof.v @@ -30,17 +30,17 @@ Record map_wf (m: mapping) : Prop := mk_map_wf { map_wf_inj: (forall id1 id2 r, - m.(map_vars)!id1 = Some r -> m.(map_vars)!id2 = Some r -> id1 = id2); + m.(map_vars)$id1 = Some r -> m.(map_vars)$id2 = Some r -> id1 = id2); map_wf_disj: (forall id r, - m.(map_vars)!id = Some r -> In r m.(map_letvars) -> False) + m.(map_vars)$id = Some r -> In r m.(map_letvars) -> False) }. Lemma init_mapping_wf: map_wf init_mapping. Proof. unfold init_mapping; split; simpl. - intros until r. rewrite PTree.gempty. congruence. + intros until r. rewrite ATree.gempty. congruence. tauto. Qed. @@ -51,8 +51,8 @@ Lemma add_var_wf: Proof. intros. monadInv H. apply mk_map_wf; simpl. - intros until r0. repeat rewrite PTree.gsspec. - destruct (peq id1 name); destruct (peq id2 name). + intros until r0. repeat rewrite ATree.gsspec. + destruct (ATree.elt_eq id1 name); destruct (ATree.elt_eq id2 name). congruence. intros. inv H. elimtype False. apply valid_fresh_absurd with r0 s1. @@ -63,8 +63,8 @@ Proof. apply H1. left; exists id1; auto. eauto with rtlg. inv H0. eauto. - intros until r0. rewrite PTree.gsspec. - destruct (peq id name). + intros until r0. rewrite ATree.gsspec. + destruct (ATree.elt_eq id name). intros. inv H. apply valid_fresh_absurd with r0 s1. apply H1. right; auto. @@ -103,7 +103,7 @@ Record match_env mk_match_env { me_vars: (forall id v, - e!id = Some v -> exists r, map.(map_vars)!id = Some r /\ Val.lessdef v rs#r); + e$id = Some v -> exists r, map.(map_vars)$id = Some r /\ Val.lessdef v rs#r); me_letvars: Val.lessdef_list le rs##(map.(map_letvars)) }. @@ -111,8 +111,8 @@ Record match_env Lemma match_env_find_var: forall map e le rs id v r, match_env map e le rs -> - e!id = Some v -> - map.(map_vars)!id = Some r -> + e$id = Some v -> + map.(map_vars)$id = Some r -> Val.lessdef v rs#r. Proof. intros. exploit me_vars; eauto. intros [r' [EQ' RS]]. @@ -176,12 +176,12 @@ Lemma match_env_update_var: forall map e le rs id r v tv, Val.lessdef v tv -> map_wf map -> - map.(map_vars)!id = Some r -> + map.(map_vars)$id = Some r -> match_env map e le rs -> - match_env map (PTree.set id v e) le (rs#r <- tv). + match_env map (ATree.set id v e) le (rs#r <- tv). Proof. intros. inversion H0. inversion H2. apply mk_match_env. - intros id' v'. rewrite PTree.gsspec. destruct (peq id' id); intros. + intros id' v'. rewrite ATree.gsspec. destruct (ATree.elt_eq id' id); intros. subst id'. inv H3. exists r; split. auto. rewrite PMap.gss. auto. exploit me_vars0; eauto. intros [r' [A B]]. exists r'; split. auto. rewrite PMap.gso; auto. @@ -249,10 +249,10 @@ Qed. Lemma match_env_empty: forall map, map.(map_letvars) = nil -> - match_env map (PTree.empty val) nil (Regmap.init Vundef). + match_env map (ATree.empty val) nil (Regmap.init Vundef). Proof. intros. apply mk_match_env. - intros. rewrite PTree.gempty in H0. discriminate. + intros. rewrite ATree.gempty in H0. discriminate. rewrite H. constructor. Qed. @@ -281,7 +281,7 @@ Proof. destruct (IHil _ _ _ _ nil nil _ EQ) as [ME UNDEF]. constructor. inv ME. split. replace (init_regs nil x) with (Regmap.init Vundef) in me_vars0, me_letvars0. constructor; simpl. - intros id v. repeat rewrite PTree.gsspec. destruct (peq id a); intros. + intros id v. repeat rewrite ATree.gsspec. destruct (ATree.elt_eq id a); intros. subst a. inv H. exists x1; split. auto. constructor. eauto. eauto. @@ -290,7 +290,7 @@ Proof. (* vl = v1 :: vs *) destruct (IHil _ _ _ _ _ _ _ EQ H0) as [ME UNDEF]. inv ME. split. constructor; simpl. - intros id v. repeat rewrite PTree.gsspec. destruct (peq id a); intros. + intros id v. repeat rewrite ATree.gsspec. destruct (ATree.elt_eq id a); intros. subst a. inv H. inv H1. exists x1; split. auto. rewrite Regmap.gss. constructor. inv H1. eexists; eauto. exploit me_vars0; eauto. intros [r' [C D]]. @@ -321,8 +321,8 @@ Proof. exploit IHil; eauto. intro. monadInv EQ1. constructor. - intros id v. simpl. repeat rewrite PTree.gsspec. - destruct (peq id a). subst a. intro. + intros id v. simpl. repeat rewrite ATree.gsspec. + destruct (ATree.elt_eq id a). subst a. intro. exists x1. split. auto. inv H3. constructor. eauto with rtlg. intros. eapply me_vars; eauto. @@ -398,7 +398,7 @@ Proof. intros until tf. unfold transl_fundef, transf_partial_fundef. case f; intro. unfold transl_function. - destruct (reserve_labels (fn_body f0) (PTree.empty node, init_state)) as [ngoto s0]. + destruct (reserve_labels (fn_body f0) (ATree.empty node, init_state)) as [ngoto s0]. case (transl_fun f0 ngoto s0); simpl; intros. discriminate. destruct p. simpl in H. inversion H. reflexivity. @@ -516,8 +516,8 @@ Definition transl_condexpr_prop corresponding to the evaluations of sub-expressions or sub-statements. *) Lemma transl_expr_Evar_correct: - forall (le : letenv) (id : positive) (v: val), - e ! id = Some v -> + forall (le : letenv) (id : ident) (v: val), + e $ id = Some v -> transl_expr_prop le (Evar id) v. Proof. intros; red; intros. inv TE. @@ -1263,7 +1263,7 @@ Qed. Lemma tr_find_label: forall c map lbl n (ngoto: labelmap) nret rret s' k' cs, - ngoto!lbl = Some n -> + ngoto$lbl = Some n -> forall s k ns1 nd1 nexits1, find_label lbl s k = Some (s', k') -> tr_stmt c map s ns1 nd1 nexits1 ngoto nret rret -> diff --git a/backend/RTLgenspec.v b/backend/RTLgenspec.v index 17022a7df5..c463bfa804 100644 --- a/backend/RTLgenspec.v +++ b/backend/RTLgenspec.v @@ -238,7 +238,7 @@ Hint Resolve regs_valid_incr: rtlg. local or let-bound variable. *) Definition reg_in_map (m: mapping) (r: reg) : Prop := - (exists id, m.(map_vars)!id = Some r) \/ In r m.(map_letvars). + (exists id, m.(map_vars)$id = Some r) \/ In r m.(map_letvars). (** A compilation environment (mapping) is valid in a given state if the registers associated with Cminor local variables and let-bound variables @@ -316,7 +316,7 @@ Lemma init_mapping_valid: Proof. unfold map_valid, init_mapping. intros s r [[id A] | B]. - simpl in A. rewrite PTree.gempty in A; discriminate. + simpl in A. rewrite ATree.gempty in A; discriminate. simpl in B. tauto. Qed. @@ -326,7 +326,7 @@ Lemma find_var_in_map: forall s1 s2 map name r i, find_var map name s1 = OK r s2 i -> reg_in_map map r. Proof. - intros until r. unfold find_var; caseEq (map.(map_vars)!name). + intros until r. unfold find_var; caseEq (ATree.get name map.(map_vars)). intros. inv H0. left; exists name; auto. intros. inv H0. Qed. @@ -371,7 +371,7 @@ Proof. intros. monadInv H. split. eauto with rtlg. inversion EQ. subst. red. intros r' [[id A] | B]. - simpl in A. rewrite PTree.gsspec in A. destruct (peq id name). + simpl in A. rewrite ATree.gsspec in A. destruct (ATree.elt_eq id name). inv A. eauto with rtlg. apply reg_valid_incr with s1. eauto with rtlg. apply H0. left; exists id; auto. @@ -381,9 +381,9 @@ Qed. Lemma add_var_find: forall s1 s2 map name r map' i, - add_var map name s1 = OK (r,map') s2 i -> map'.(map_vars)!name = Some r. + add_var map name s1 = OK (r,map') s2 i -> map'.(map_vars)$name = Some r. Proof. - intros. monadInv H. simpl. apply PTree.gss. + intros. monadInv H. simpl. apply ATree.gss. Qed. Lemma add_vars_valid: @@ -517,7 +517,7 @@ Qed. Inductive target_reg_ok (map: mapping) (pr: list reg): expr -> reg -> Prop := | target_reg_var: forall id r, - map.(map_vars)!id = Some r -> + map.(map_vars)$id = Some r -> target_reg_ok map pr (Evar id) r | target_reg_letvar: forall idx r, @@ -587,7 +587,7 @@ Proof. intros. unfold alloc_reg in H1. destruct a; try (eapply new_reg_target_ok; eauto; fail). (* Evar *) - generalize H1; unfold find_var. caseEq (map_vars map)!i0; intros. + generalize H1; unfold find_var. caseEq (map_vars map)$i0; intros. inv H3. constructor. auto. inv H3. (* Elet *) generalize H1; unfold find_letvar. caseEq (nth_error (map_letvars map) n); intros. @@ -673,7 +673,7 @@ Inductive reg_map_ok: mapping -> reg -> option ident -> Prop := ~reg_in_map map rd -> reg_map_ok map rd None | reg_map_ok_somevar: forall map rd id, - map.(map_vars)!id = Some rd -> + map.(map_vars)$id = Some rd -> reg_map_ok map rd (Some id). Hint Resolve reg_map_ok_novar: rtlg. @@ -696,7 +696,7 @@ Hint Resolve reg_map_ok_novar: rtlg. Inductive tr_expr (c: code): mapping -> list reg -> expr -> node -> node -> reg -> option ident -> Prop := | tr_Evar: forall map pr id ns nd r rd dst, - map.(map_vars)!id = Some r -> + map.(map_vars)$id = Some r -> ((rd = r /\ dst = None) \/ (reg_map_ok map rd dst /\ ~In rd pr)) -> tr_move c ns r nd rd -> tr_expr c map pr (Evar id) ns nd rd dst @@ -811,7 +811,7 @@ Inductive tr_exitexpr (c: code): Inductive tr_builtin_res: mapping -> builtin_res ident -> builtin_res reg -> Prop := | tr_builtin_res_var: forall map id r, - map.(map_vars)!id = Some r -> + map.(map_vars)$id = Some r -> tr_builtin_res map (BR id) (BR r) | tr_builtin_res_none: forall map, tr_builtin_res map BR_none BR_none @@ -833,7 +833,7 @@ Inductive tr_stmt (c: code) (map: mapping): | tr_Sskip: forall ns nexits ngoto nret rret, tr_stmt c map Sskip ns ns nexits ngoto nret rret | tr_Sassign: forall id a ns nd nexits ngoto nret rret r, - map.(map_vars)!id = Some r -> + map.(map_vars)$id = Some r -> tr_expr c map nil a ns nd r (Some id) -> tr_stmt c map (Sassign id a) ns nd nexits ngoto nret rret | tr_Sstore: forall chunk addr al b ns nd nexits ngoto nret rret rd n1 rl n2, @@ -895,12 +895,12 @@ Inductive tr_stmt (c: code) (map: mapping): tr_expr c map nil a ns nret rret None -> tr_stmt c map (Sreturn (Some a)) ns nd nexits ngoto nret (Some rret) | tr_Slabel: forall lbl s ns nd nexits ngoto nret rret n, - ngoto!lbl = Some n -> + ngoto$lbl = Some n -> c!n = Some (Inop ns) -> tr_stmt c map s ns nd nexits ngoto nret rret -> tr_stmt c map (Slabel lbl s) ns nd nexits ngoto nret rret | tr_Sgoto: forall lbl ns nd nexits ngoto nret rret, - ngoto!lbl = Some ns -> + ngoto$lbl = Some ns -> tr_stmt c map (Sgoto lbl) ns nd nexits ngoto nret rret. (** [tr_function f tf] specifies the RTL function [tf] that @@ -1002,7 +1002,7 @@ with transl_condexpr_charact: Proof. induction a; intros; try (monadInv TR); saturateTrans. (* Evar *) - generalize EQ; unfold find_var. caseEq (map_vars map)!i; intros; inv EQ1. + generalize EQ; unfold find_var. caseEq (map_vars map)$i; intros; inv EQ1. econstructor; eauto. inv OK. left; split; congruence. right; eauto with rtlg. eapply add_move_charact; eauto. @@ -1096,7 +1096,7 @@ Lemma transl_expr_assign_charact: Proof. induction a; intros; monadInv TR; saturateTrans. (* Evar *) - generalize EQ; unfold find_var. caseEq (map_vars map)!i; intros; inv EQ1. + generalize EQ; unfold find_var. caseEq (map_vars map)$i; intros; inv EQ1. econstructor; eauto. eapply add_move_charact; eauto. (* Eop *) @@ -1140,7 +1140,7 @@ Lemma alloc_optreg_map_ok: reg_map_ok map r optid. Proof. unfold alloc_optreg; intros. destruct optid. - constructor. unfold find_var in H0. destruct (map_vars map)!i0; monadInv H0. auto. + constructor. unfold find_var in H0. destruct (map_vars map)$i0; monadInv H0. auto. constructor. eapply new_reg_not_in_map; eauto. Qed. @@ -1223,7 +1223,7 @@ Lemma convert_builtin_res_charact: tr_builtin_res map res res'. Proof. destruct res; simpl; intros. -- monadInv TR. constructor. unfold find_var in EQ. destruct (map_vars map)!x; inv EQ; auto. +- monadInv TR. constructor. unfold find_var in EQ. destruct (map_vars map)$x; inv EQ; auto. - destruct oty; monadInv TR. + constructor. eauto with rtlg. + constructor. @@ -1241,7 +1241,7 @@ Proof. (* Sskip *) constructor. (* Sassign *) - revert EQ. unfold find_var. case_eq (map_vars map)!i; intros; monadInv EQ. + revert EQ. unfold find_var. case_eq (map_vars map)$i; intros; monadInv EQ. eapply tr_Sassign; eauto. eapply transl_expr_assign_charact; eauto with rtlg. constructor. auto. @@ -1323,13 +1323,13 @@ Proof. constructor. auto. simpl; tauto. monadInv TR. constructor. (* Slabel *) - generalize EQ0; clear EQ0. case_eq (ngoto!l); intros; monadInv EQ0. + generalize EQ0; clear EQ0. case_eq (ngoto$l); intros; monadInv EQ0. generalize EQ1; clear EQ1. unfold handle_error. case_eq (update_instr n (Inop ns) s0); intros; inv EQ1. econstructor. eauto. eauto with rtlg. eapply tr_stmt_incr with s0; eauto with rtlg. (* Sgoto *) - generalize TR; clear TR. case_eq (ngoto!l); intros; monadInv TR. + generalize TR; clear TR. case_eq (ngoto$l); intros; monadInv TR. econstructor. auto. Qed. @@ -1339,7 +1339,7 @@ Lemma transl_function_charact: tr_function f tf. Proof. intros until tf. unfold transl_function. - caseEq (reserve_labels (fn_body f) (PTree.empty node, init_state)). + caseEq (reserve_labels (fn_body f) (ATree.empty node, init_state)). intros ngoto s0 RESERVE. caseEq (transl_fun f ngoto s0). congruence. intros [nentry rparams] sfinal INCR TR E. inv E. diff --git a/backend/Selection.v b/backend/Selection.v index 4520cb0c54..c06bdace81 100644 --- a/backend/Selection.v +++ b/backend/Selection.v @@ -62,7 +62,7 @@ Definition store (chunk: memory_chunk) (e1 e2: expr) := Section SELECTION. Definition globdef := AST.globdef Cminor.fundef unit. -Variable defmap: PTree.t globdef. +Variable defmap: ATree.t globdef. Context {hf: helper_functions}. Definition sel_constant (cst: Cminor.constant) : expr := @@ -192,7 +192,7 @@ Definition classify_call (e: Cminor.expr) : call_kind := match expr_is_addrof_ident e with | None => Call_default | Some id => - match defmap!id with + match defmap$id with | Some(Gfun(External ef)) => if ef_inline ef then Call_builtin ef else Call_imm id | _ => Call_imm id end @@ -322,7 +322,7 @@ End SELECTION. (** Conversion of functions. *) -Definition sel_function (dm: PTree.t globdef) (hf: helper_functions) (f: Cminor.function) : res function := +Definition sel_function (dm: ATree.t globdef) (hf: helper_functions) (f: Cminor.function) : res function := do body' <- sel_stmt dm f.(Cminor.fn_body); OK (mkfunction f.(Cminor.fn_sig) @@ -331,7 +331,7 @@ Definition sel_function (dm: PTree.t globdef) (hf: helper_functions) (f: Cminor. f.(Cminor.fn_stackspace) body'). -Definition sel_fundef (dm: PTree.t globdef) (hf: helper_functions) (f: Cminor.fundef) : res fundef := +Definition sel_fundef (dm: ATree.t globdef) (hf: helper_functions) (f: Cminor.fundef) : res fundef := transf_partial_fundef (sel_function dm hf) f. (** Setting up the helper functions. *) @@ -349,10 +349,10 @@ Definition globdef_of_interest (gd: globdef) : bool := | _ => false end. -Definition record_globdefs (defmap: PTree.t globdef) : PTree.t globdef := - PTree.fold - (fun m id gd => if globdef_of_interest gd then PTree.set id gd m else m) - defmap (PTree.empty globdef). +Definition record_globdefs (defmap: ATree.t globdef) : ATree.t globdef := + ATree.fold + (fun m id gd => if globdef_of_interest gd then ATree.set id gd m else m) + defmap (ATree.empty globdef). Definition lookup_helper_aux (name: String.string) (sg: signature) (res: option ident) @@ -365,16 +365,16 @@ Definition lookup_helper_aux | _ => res end. -Definition lookup_helper (globs: PTree.t globdef) +Definition lookup_helper (globs: ATree.t globdef) (name: String.string) (sg: signature) : res ident := - match PTree.fold (lookup_helper_aux name sg) globs None with + match ATree.fold (lookup_helper_aux name sg) globs None with | Some id => OK id | None => Error (MSG name :: MSG ": missing or incorrect declaration" :: nil) end. Local Open Scope string_scope. -Definition get_helpers (defmap: PTree.t globdef) : res helper_functions := +Definition get_helpers (defmap: ATree.t globdef) : res helper_functions := let globs := record_globdefs defmap in do i64_dtos <- lookup_helper globs "__compcert_i64_dtos" sig_f_l ; do i64_dtou <- lookup_helper globs "__compcert_i64_dtou" sig_f_l ; diff --git a/backend/Selectionproof.v b/backend/Selectionproof.v index dc01ad20b3..3de8f75b8b 100644 --- a/backend/Selectionproof.v +++ b/backend/Selectionproof.v @@ -33,31 +33,31 @@ Definition match_prog (p: Cminor.program) (tp: CminorSel.program) := (** Processing of helper functions *) Lemma record_globdefs_sound: - forall dm id gd, (record_globdefs dm)!id = Some gd -> dm!id = Some gd. + forall dm id gd, (record_globdefs dm)$id = Some gd -> dm$id = Some gd. Proof. intros. - set (f := fun m id gd => if globdef_of_interest gd then PTree.set id gd m else m) in *. - set (P := fun m m' => m'!id = Some gd -> m!id = Some gd). - assert (X: P dm (PTree.fold f dm (PTree.empty _))). - { apply PTree_Properties.fold_rec. + set (f := fun m id gd => if globdef_of_interest gd then ATree.set id gd m else m) in *. + set (P := fun m m' => m'$id = Some gd -> m$id = Some gd). + assert (X: P dm (ATree.fold f dm (ATree.empty _))). + { apply ATree_Properties.fold_rec. - unfold P; intros. rewrite <- H0; auto. - - red. rewrite ! PTree.gempty. auto. - - unfold P; intros. rewrite PTree.gsspec. unfold f in H3. + - red. rewrite ! ATree.gempty. auto. + - unfold P; intros. rewrite ATree.gsspec. unfold f in H3. destruct (globdef_of_interest v). - + rewrite PTree.gsspec in H3. destruct (peq id k); auto. - + apply H2 in H3. destruct (peq id k). congruence. auto. } + + rewrite ATree.gsspec in H3. destruct (ATree.elt_eq id k); auto. + + apply H2 in H3. destruct (ATree.elt_eq id k). congruence. auto. } apply X. auto. Qed. Lemma lookup_helper_correct_1: forall globs name sg id, lookup_helper globs name sg = OK id -> - globs!id = Some (Gfun (External (EF_runtime name sg))). + globs$id = Some (Gfun (External (EF_runtime name sg))). Proof. intros. - set (P := fun (m: PTree.t globdef) res => res = Some id -> m!id = Some(Gfun(External (EF_runtime name sg)))). - assert (P globs (PTree.fold (lookup_helper_aux name sg) globs None)). - { apply PTree_Properties.fold_rec; red; intros. + set (P := fun (m: ATree.t globdef) res => res = Some id -> m$id = Some(Gfun(External (EF_runtime name sg)))). + assert (P globs (ATree.fold (lookup_helper_aux name sg) globs None)). + { apply ATree_Properties.fold_rec; red; intros. - rewrite <- H0. apply H1; auto. - discriminate. - assert (EITHER: k = id /\ v = Gfun (External (EF_runtime name sg)) @@ -67,11 +67,11 @@ Proof. destruct (signature_eq sg sg0); auto. inversion H3. left; split; auto. repeat f_equal; auto. } destruct EITHER as [[X Y] | X]. - subst k v. apply PTree.gss. - apply H2 in X. rewrite PTree.gso by congruence. auto. + subst k v. apply ATree.gss. + apply H2 in X. rewrite ATree.gso by congruence. auto. } red in H0. unfold lookup_helper in H. - destruct (PTree.fold (lookup_helper_aux name sg) globs None); inv H. auto. + destruct (ATree.fold (lookup_helper_aux name sg) globs None); inv H. auto. Qed. Lemma lookup_helper_correct: @@ -357,7 +357,7 @@ Proof. destruct (Genv.find_symbol ge id) as [b|] eqn:FS; try discriminate. rewrite Genv.find_funct_find_funct_ptr in H1. assert (DFL: exists b1, Genv.find_symbol ge id = Some b1 /\ Vptr b Ptrofs.zero = Vptr b1 Ptrofs.zero) by (exists b; auto). - unfold globdef; destruct (prog_defmap unit)!id as [[[f|ef] |gv] |] eqn:G; auto. + unfold globdef; destruct (prog_defmap unit)$id as [[[f|ef] |gv] |] eqn:G; auto. destruct (ef_inline ef) eqn:INLINE; auto. destruct (prog_defmap_linkorder _ _ _ _ H G) as (gd & P & Q). inv Q. inv H2. @@ -601,14 +601,14 @@ Qed. (** Relationship between the local environments. *) Definition env_lessdef (e1 e2: env) : Prop := - forall id v1, e1!id = Some v1 -> exists v2, e2!id = Some v2 /\ Val.lessdef v1 v2. + forall id v1, e1$id = Some v1 -> exists v2, e2$id = Some v2 /\ Val.lessdef v1 v2. Lemma set_var_lessdef: forall e1 e2 id v1 v2, env_lessdef e1 e2 -> Val.lessdef v1 v2 -> - env_lessdef (PTree.set id v1 e1) (PTree.set id v2 e2). + env_lessdef (ATree.set id v1 e1) (ATree.set id v2 e2). Proof. - intros; red; intros. rewrite PTree.gsspec in *. destruct (peq id0 id). + intros; red; intros. rewrite ATree.gsspec in *. destruct (ATree.elt_eq id0 id). exists v2; split; congruence. auto. Qed. @@ -627,7 +627,7 @@ Lemma set_params_lessdef: env_lessdef (set_params vl1 il) (set_params vl2 il). Proof. induction il; simpl; intros. - red; intros. rewrite PTree.gempty in H0; congruence. + red; intros. rewrite ATree.gempty in H0; congruence. inv H; apply set_var_lessdef; auto. Qed. @@ -837,9 +837,9 @@ Remark match_call_cont_cont: forall k k', match_call_cont k k' -> exists cunit hf, match_cont cunit hf k k'. Proof. intros. simple refine (let cunit : Cminor.program := _ in _). - econstructor. apply nil. apply nil. apply xH. + econstructor. apply nil. apply nil. apply #"!dummy_main!". simple refine (let hf : helper_functions := _ in _). - econstructor; apply xH. + econstructor; apply #"!dummy_helper!". exists cunit, hf; auto. Qed. diff --git a/backend/SplitLongproof.v b/backend/SplitLongproof.v index f1e8b5904b..0239a3db8a 100644 --- a/backend/SplitLongproof.v +++ b/backend/SplitLongproof.v @@ -53,7 +53,7 @@ Axiom i64_helpers_correct : /\ (forall x y, external_implements "__compcert_i64_smulh" sig_ll_l (x::y::nil) (Val.mullhs x y)). Definition helper_declared {F V: Type} (p: AST.program (AST.fundef F) V) (id: ident) (name: string) (sg: signature) : Prop := - (prog_defmap p)!id = Some (Gfun (External (EF_runtime name sg))). + (prog_defmap p)$id = Some (Gfun (External (EF_runtime name sg))). Definition helper_functions_declared {F V: Type} (p: AST.program (AST.fundef F) V) (hf: helper_functions) : Prop := helper_declared p i64_dtos "__compcert_i64_dtos" sig_f_l diff --git a/backend/Unusedglob.v b/backend/Unusedglob.v index 916e111b06..2d064d8fa4 100644 --- a/backend/Unusedglob.v +++ b/backend/Unusedglob.v @@ -22,7 +22,7 @@ Local Open Scope string_scope. (** The working set *) -Module IS := FSetAVL.Make(OrderedPositive). +Module IS := FSetAVL.Make(OrderedIdent). Record workset := mk_workset { w_seen :> IS.t; @@ -73,10 +73,10 @@ Definition add_ref_init_data (w: workset) (i: init_data) : workset := Definition add_ref_globvar (gv: globvar unit) (w: workset): workset := List.fold_left add_ref_init_data gv.(gvar_init) w. -Definition prog_map : Type := PTree.t (globdef fundef unit). +Definition prog_map : Type := ATree.t (globdef fundef unit). Definition add_ref_definition (pm: prog_map) (id: ident) (w: workset): workset := - match pm!id with + match pm$id with | None => w | Some (Gfun (Internal f)) => add_ref_function f w | Some (Gfun (External ef)) => w @@ -124,7 +124,7 @@ Fixpoint filter_globdefs (used: IS.t) (accu defs: list (ident * globdef fundef u the possible exception of the [prog_main] name. *) Definition global_defined (p: program) (pm: prog_map) (id: ident) : bool := - match pm!id with Some _ => true | None => ident_eq id (prog_main p) end. + match pm$id with Some _ => true | None => ident_eq id (prog_main p) end. Definition transform_program (p: program) : res program := let pm := prog_defmap p in diff --git a/backend/Unusedglobproof.v b/backend/Unusedglobproof.v index 7899a04cc9..ba6bd742de 100644 --- a/backend/Unusedglobproof.v +++ b/backend/Unusedglobproof.v @@ -34,7 +34,7 @@ Record match_prog_1 (u: IS.t) (p tp: program) : Prop := { tp.(prog_public) = p.(prog_public); match_prog_def: forall id, - (prog_defmap tp)!id = if IS.mem id u then (prog_defmap p)!id else None; + (prog_defmap tp)$id = if IS.mem id u then (prog_defmap p)$id else None; match_prog_unique: list_norepet (prog_defs_names tp) }. @@ -59,7 +59,7 @@ Definition ref_def (gd: globdef fundef unit) (id: ident) : Prop := Record valid_used_set (p: program) (u: IS.t) : Prop := { used_closed: forall id gd id', - IS.In id u -> (prog_defmap p)!id = Some gd -> ref_def gd id' -> + IS.In id u -> (prog_defmap p)$id = Some gd -> ref_def gd id' -> IS.In id' u; used_main: IS.In p.(prog_main) u; @@ -148,7 +148,7 @@ Lemma add_ref_definition_incl: forall pm id w, workset_incl w (add_ref_definition pm id w). Proof. unfold add_ref_definition; intros. - destruct (pm!id) as [[[] | ? ] | ]. + destruct (pm$id) as [[[] | ? ] | ]. apply add_ref_function_incl. apply workset_incl_refl. apply add_ref_globvar_incl. @@ -202,7 +202,7 @@ Qed. Lemma seen_add_ref_definition: forall pm id gd id' w, - pm!id = Some gd -> ref_def gd id' -> IS.In id' (add_ref_definition pm id w). + pm$id = Some gd -> ref_def gd id' -> IS.In id' (add_ref_definition pm id w). Proof. unfold add_ref_definition; intros. rewrite H. red in H0; destruct gd as [[f|ef]|gv]. apply seen_add_ref_function; auto. @@ -255,12 +255,12 @@ Let pm := prog_defmap p. Definition workset_invariant (w: workset) : Prop := forall id gd id', - IS.In id w -> ~List.In id (w_todo w) -> pm!id = Some gd -> ref_def gd id' -> + IS.In id w -> ~List.In id (w_todo w) -> pm$id = Some gd -> ref_def gd id' -> IS.In id' w. Definition used_set_closed (u: IS.t) : Prop := forall id gd id', - IS.In id u -> pm!id = Some gd -> ref_def gd id' -> IS.In id' u. + IS.In id u -> pm$id = Some gd -> ref_def gd id' -> IS.In id' u. Lemma iter_step_invariant: forall w, @@ -320,7 +320,7 @@ Proof. - intros. eapply used_globals_incl; eauto. apply seen_public_initial_workset; auto. - intros. apply ISF.for_all_iff in H0. + red in H0. apply H0 in H1. unfold global_defined in H1. - destruct pm!id as [g|] eqn:E. + destruct pm$id as [g|] eqn:E. * left. change id with (fst (id,g)). apply in_map. apply in_prog_defmap; auto. * InvBooleans; auto. + hnf. simpl; intros; congruence. @@ -335,7 +335,7 @@ Section TRANSFORMATION. Variable p: program. Variable used: IS.t. -Let add_def (m: prog_map) idg := PTree.set (fst idg) (snd idg) m. +Let add_def (m: prog_map) idg := ATree.set (fst idg) (snd idg) m. Remark filter_globdefs_accu: forall defs accu1 accu2 u, @@ -357,14 +357,14 @@ Qed. Lemma filter_globdefs_map_1: forall id l u m1, IS.mem id u = false -> - m1!id = None -> - (fold_left add_def (filter_globdefs u nil l) m1)!id = None. + m1$id = None -> + (fold_left add_def (filter_globdefs u nil l) m1)$id = None. Proof. induction l as [ | [id1 gd1] l]; simpl; intros. - auto. - destruct (IS.mem id1 u) eqn:MEM. + rewrite filter_globdefs_nil. rewrite fold_left_app. simpl. - unfold add_def at 1. simpl. rewrite PTree.gso by congruence. eapply IHl; eauto. + unfold add_def at 1. simpl. rewrite ATree.gso by congruence. eapply IHl; eauto. rewrite ISF.remove_b. rewrite H; auto. + eapply IHl; eauto. Qed. @@ -372,8 +372,8 @@ Qed. Lemma filter_globdefs_map_2: forall id l u m1 m2, IS.mem id u = true -> - m1!id = m2!id -> - (fold_left add_def (filter_globdefs u nil l) m1)!id = (fold_left add_def (List.rev l) m2)!id. + m1$id = m2$id -> + (fold_left add_def (filter_globdefs u nil l) m1)$id = (fold_left add_def (List.rev l) m2)$id. Proof. induction l as [ | [id1 gd1] l]; simpl; intros. - auto. @@ -381,22 +381,22 @@ Proof. destruct (IS.mem id1 u) eqn:MEM. + rewrite filter_globdefs_nil. rewrite fold_left_app. simpl. unfold add_def at 1 3. simpl. - rewrite ! PTree.gsspec. destruct (peq id id1). auto. + rewrite ! ATree.gsspec. destruct (ATree.elt_eq id id1). auto. apply IHl; auto. apply IS.mem_1. apply IS.remove_2; auto. apply IS.mem_2; auto. -+ unfold add_def at 2. simpl. rewrite PTree.gso by congruence. apply IHl; auto. ++ unfold add_def at 2. simpl. rewrite ATree.gso by congruence. apply IHl; auto. Qed. Lemma filter_globdefs_map: forall id u defs, - (PTree_Properties.of_list (filter_globdefs u nil (List.rev defs)))! id = - if IS.mem id u then (PTree_Properties.of_list defs)!id else None. + (ATree_Properties.of_list (filter_globdefs u nil (List.rev defs)))$ id = + if IS.mem id u then (ATree_Properties.of_list defs)$id else None. Proof. - intros. unfold PTree_Properties.of_list. fold prog_map. unfold PTree.elt. fold add_def. + intros. unfold ATree_Properties.of_list. fold prog_map. unfold ATree.elt. fold add_def. destruct (IS.mem id u) eqn:MEM. - erewrite filter_globdefs_map_2. rewrite List.rev_involutive. reflexivity. auto. auto. -- apply filter_globdefs_map_1. auto. apply PTree.gempty. +- apply filter_globdefs_map_1. auto. apply ATree.gempty. Qed. Lemma filter_globdefs_domain: @@ -457,7 +457,7 @@ Definition kept (id: ident) : Prop := IS.In id used. Lemma kept_closed: forall id gd id', - kept id -> pm!id = Some gd -> ref_def gd id' -> kept id'. + kept id -> pm$id = Some gd -> ref_def gd id' -> kept id'. Proof. intros. eapply used_closed; eauto. Qed. @@ -481,7 +481,7 @@ Lemma transform_find_symbol_1: Genv.find_symbol ge id = Some b -> kept id -> exists b', Genv.find_symbol tge id = Some b'. Proof. intros. - assert (A: exists g, (prog_defmap p)!id = Some g). + assert (A: exists g, (prog_defmap p)$id = Some g). { apply prog_defmap_dom. eapply Genv.find_symbol_inversion; eauto. } destruct A as (g & P). apply Genv.find_symbol_exists with g. @@ -494,7 +494,7 @@ Lemma transform_find_symbol_2: Genv.find_symbol tge id = Some b -> kept id /\ exists b', Genv.find_symbol ge id = Some b'. Proof. intros. - assert (A: exists g, (prog_defmap tp)!id = Some g). + assert (A: exists g, (prog_defmap tp)$id = Some g). { apply prog_defmap_dom. eapply Genv.find_symbol_inversion; eauto. } destruct A as (g & P). erewrite match_prog_def in P by eauto. @@ -568,15 +568,15 @@ Proof. exists b; split; auto. eapply init_meminj_eq; eauto. - exploit init_meminj_invert; eauto. intros (A & id & B & C). assert (kept id) by (eapply transform_find_symbol_2; eauto). - assert (pm!id = Some gd). + assert (pm$id = Some gd). { unfold pm; rewrite Genv.find_def_symbol. exists b; auto. } - assert ((prog_defmap tp)!id = Some gd). + assert ((prog_defmap tp)$id = Some gd). { erewrite match_prog_def by eauto. rewrite IS.mem_1 by auto. auto. } rewrite Genv.find_def_symbol in H3. destruct H3 as (b1 & P & Q). fold tge in P. replace b' with b1 by congruence. split; auto. split; auto. intros. eapply kept_closed; eauto. - exploit init_meminj_invert; eauto. intros (A & id & B & C). - assert ((prog_defmap tp)!id = Some gd). + assert ((prog_defmap tp)$id = Some gd). { rewrite Genv.find_def_symbol. exists b'; auto. } erewrite match_prog_def in H1 by eauto. destruct (IS.mem id used); try discriminate. @@ -1079,7 +1079,7 @@ Lemma init_meminj_invert_strong: /\ (forall i, ref_def gd i -> kept i). Proof. intros. exploit init_meminj_invert; eauto. intros (A & id & B & C). - assert (exists gd, (prog_defmap p)!id = Some gd). + assert (exists gd, (prog_defmap p)$id = Some gd). { apply prog_defmap_dom. eapply Genv.find_symbol_inversion; eauto. } destruct H0 as [gd DM]. rewrite Genv.find_def_symbol in DM. destruct DM as (b'' & P & Q). fold ge in P. rewrite P in B; inv B. @@ -1203,7 +1203,7 @@ Lemma init_mem_exists: Proof. intros. apply Genv.init_mem_exists. intros. - assert (P: (prog_defmap tp)!id = Some (Gvar v)). + assert (P: (prog_defmap tp)$id = Some (Gvar v)). { eapply prog_defmap_norepet; eauto. eapply match_prog_unique; eauto. } rewrite (match_prog_def _ _ _ TRANSF) in P. destruct (IS.mem id used) eqn:U; try discriminate. exploit Genv.init_mem_inversion; eauto. apply in_prog_defmap; eauto. intros [AL FV]. @@ -1305,7 +1305,7 @@ Qed. Remark used_not_defined: forall p used id, valid_used_set p used -> - (prog_defmap p)!id = None -> + (prog_defmap p)$id = None -> IS.mem id used = false \/ id = prog_main p. Proof. intros. destruct (IS.mem id used) eqn:M; auto. @@ -1317,7 +1317,7 @@ Remark used_not_defined_2: forall p used id, valid_used_set p used -> id <> prog_main p -> - (prog_defmap p)!id = None -> + (prog_defmap p)$id = None -> ~IS.In id used. Proof. intros. exploit used_not_defined; eauto. intros [A|A]. @@ -1336,9 +1336,9 @@ Proof. destruct (link_prog_inv _ _ _ L) as (X & Y & Z). rewrite Z; clear Z; constructor. - intros. rewrite ISF.union_iff in H. rewrite ISF.union_iff. - rewrite prog_defmap_elements, PTree.gcombine in H0. - destruct (prog_defmap p1)!id as [gd1|] eqn:GD1; - destruct (prog_defmap p2)!id as [gd2|] eqn:GD2; + rewrite prog_defmap_elements, ATree.gcombine in H0. + destruct (prog_defmap p1)$id as [gd1|] eqn:GD1; + destruct (prog_defmap p2)$id as [gd2|] eqn:GD2; simpl in H0; try discriminate. + (* common definition *) exploit Y; eauto. intros (PUB1 & PUB2 & _). @@ -1367,20 +1367,19 @@ Proof. - intros. rewrite ISF.union_iff in H. destruct (ident_eq id (prog_main p1)). + right; assumption. -+ assert (E: exists g, link_prog_merge (prog_defmap p1)!id (prog_defmap p2)!id = Some g). - { destruct (prog_defmap p1)!id as [gd1|] eqn:GD1; - destruct (prog_defmap p2)!id as [gd2|] eqn:GD2; simpl. ++ assert (E: exists g, link_prog_merge (prog_defmap p1)$id (prog_defmap p2)$id = Some g). + { destruct (prog_defmap p1)$id as [gd1|] eqn:GD1; + destruct (prog_defmap p2)$id as [gd2|] eqn:GD2; simpl. * apply Y with id; auto. * exists gd1; auto. * exists gd2; auto. * eapply used_not_defined_2 in GD1; eauto. eapply used_not_defined_2 in GD2; eauto. tauto. - congruence. } destruct E as [g LD]. left. unfold prog_defs_names; simpl. - change id with (fst (id, g)). apply in_map. apply PTree.elements_correct. - rewrite PTree.gcombine; auto. + change id with (fst (id, g)). apply in_map. apply ATree.elements_correct. + rewrite ATree.gcombine; auto. Qed. Theorem link_match_program: @@ -1408,12 +1407,12 @@ Proof. + rewrite W. constructor; simpl; intros. * eapply match_prog_main; eauto. * rewrite (match_prog_public _ _ _ B1), (match_prog_public _ _ _ B2). auto. -* rewrite ! prog_defmap_elements, !PTree.gcombine by auto. +* rewrite ! prog_defmap_elements, !ATree.gcombine by auto. rewrite (match_prog_def _ _ _ B1 id), (match_prog_def _ _ _ B2 id). rewrite ISF.union_b. { - destruct (prog_defmap p1)!id as [gd1|] eqn:GD1; - destruct (prog_defmap p2)!id as [gd2|] eqn:GD2. + destruct (prog_defmap p1)$id as [gd1|] eqn:GD1; + destruct (prog_defmap p2)$id as [gd2|] eqn:GD2. - (* both defined *) exploit V; eauto. intros (PUB1 & PUB2 & _). assert (EQ1: IS.mem id used1 = true) by (apply IS.mem_1; eapply used_public; eauto). @@ -1432,7 +1431,7 @@ Proof. - (* none defined *) destruct (IS.mem id used1), (IS.mem id used2); auto. } -* intros. apply PTree.elements_keys_norepet. +* intros. apply ATree.elements_keys_norepet. Qed. Instance TransfSelectionLink : TransfLink match_prog := link_match_program. diff --git a/backend/ValueAnalysis.v b/backend/ValueAnalysis.v index 1f80c293b1..2ca154c0cc 100644 --- a/backend/ValueAnalysis.v +++ b/backend/ValueAnalysis.v @@ -32,7 +32,7 @@ Definition mafter_public_call : amem := mtop. Definition mafter_private_call (am_before: amem) : amem := {| am_stack := am_before.(am_stack); - am_glob := PTree.empty _; + am_glob := ATree.empty _; am_nonstack := Nonstack; am_top := plub (ab_summary (am_stack am_before)) Nonstack |}. @@ -167,7 +167,7 @@ Module DS := Dataflow_Solver(VA)(NodeSetForward). Definition mfunction_entry := {| am_stack := ablock_init Pbot; - am_glob := PTree.empty _; + am_glob := ATree.empty _; am_nonstack := Nonstack; am_top := Nonstack |}. @@ -224,15 +224,15 @@ Definition definitive_initializer (init: list init_data) : bool := Definition alloc_global (rm: romem) (idg: ident * globdef fundef unit): romem := match idg with | (id, Gfun f) => - PTree.remove id rm + ATree.remove id rm | (id, Gvar v) => if v.(gvar_readonly) && negb v.(gvar_volatile) && definitive_initializer v.(gvar_init) - then PTree.set id (store_init_data_list (ablock_init Pbot) 0 v.(gvar_init)) rm - else PTree.remove id rm + then ATree.set id (store_init_data_list (ablock_init Pbot) 0 v.(gvar_init)) rm + else ATree.remove id rm end. Definition romem_for (p: program) : romem := - List.fold_left alloc_global p.(prog_defs) (PTree.empty _). + List.fold_left alloc_global p.(prog_defs) (ATree.empty _). (** * Soundness proof *) @@ -400,11 +400,11 @@ Proof. Qed. Lemma mmatch_no_stack: forall m am astk, - mmatch bc m am -> mmatch bc m {| am_stack := astk; am_glob := PTree.empty _; am_nonstack := Nonstack; am_top := Nonstack |}. + mmatch bc m am -> mmatch bc m {| am_stack := astk; am_glob := ATree.empty _; am_nonstack := Nonstack; am_top := Nonstack |}. Proof. intros. destruct H. constructor; simpl; intros. - elim (NOSTACK b); auto. -- rewrite PTree.gempty in H0; discriminate. +- rewrite ATree.gempty in H0; discriminate. - eapply smatch_no_stack; eauto. - eapply smatch_no_stack; eauto. - auto. @@ -492,7 +492,7 @@ Proof. subst b. apply SMSTACK. elim (NOSTACK b); auto. + (* globals *) - rewrite PTree.gempty in H0; discriminate. + rewrite ATree.gempty in H0; discriminate. + (* nonstack *) destruct (eq_block b sp). congruence. eapply SM; auto. eapply mmatch_nonstack; eauto. + (* top *) @@ -582,7 +582,7 @@ Proof. - (* mmatch top *) constructor; simpl; intros. + destruct (eq_block b sp). congruence. elim n. eapply bc_stack; eauto. - + rewrite PTree.gempty in H0; discriminate. + + rewrite ATree.gempty in H0; discriminate. + destruct (eq_block b sp). subst b. eapply SM. eapply mmatch_stack; eauto. eapply SM. eapply mmatch_nonstack; eauto. @@ -669,7 +669,7 @@ Proof. - (* mmatch top *) constructor; simpl; intros. + destruct (eq_block b sp). congruence. elim n. eapply bc_stack; eauto. - + rewrite PTree.gempty in H0; discriminate. + + rewrite ATree.gempty in H0; discriminate. + destruct (eq_block b sp). congruence. eapply SM. eauto. eapply mmatch_nonstack; eauto. + destruct (eq_block b sp). congruence. @@ -756,7 +756,7 @@ Proof. subst b. eapply SM. eapply mmatch_nonstack; eauto. congruence. elim (NOSTACK b); auto. + (* globals *) - rewrite PTree.gempty in H0; discriminate. + rewrite ATree.gempty in H0; discriminate. + (* nonstack *) destruct (eq_block b sp). congruence. eapply SM; auto. eapply mmatch_nonstack; eauto. + (* top *) @@ -867,7 +867,7 @@ Proof. subst b. exact BSTK. elim (NOSTACK b); auto. + (* globals *) - rewrite PTree.gempty in H0; discriminate. + rewrite ATree.gempty in H0; discriminate. + (* nonstack *) destruct (eq_block b sp). congruence. eapply SM; auto. eapply mmatch_nonstack; eauto. + (* top *) @@ -1005,7 +1005,7 @@ Proof. - (* mmatch top *) constructor; simpl; intros. + apply ablock_init_sound. apply SMTOP. simpl; congruence. - + rewrite PTree.gempty in H0; discriminate. + + rewrite ATree.gempty in H0; discriminate. + apply SMTOP; auto. + apply SMTOP; auto. + red; simpl; intros. destruct (plt b (Mem.nextblock m)). @@ -1692,7 +1692,7 @@ Proof. - destruct (Mem.alloc m 0 1) as [m1 b1] eqn:ALLOC. unfold Genv.find_symbol in H2; simpl in H2. unfold Genv.find_var_info, Genv.find_def in H3; simpl in H3. - rewrite PTree.gsspec in H2. destruct (peq id id1). + rewrite ATree.gsspec in H2. destruct (ATree.elt_eq id id1). inv H2. rewrite PTree.gss in H3. discriminate. assert (Plt b (Genv.genv_next g)) by (eapply Genv.genv_symb_range; eauto). rewrite PTree.gso in H3 by (apply Plt_ne; auto). @@ -1709,7 +1709,7 @@ Proof. destruct (Genv.store_init_data_list ge m2 b1 0 (gvar_init gv)) as [m3 | ] eqn:SIDL; try discriminate. unfold Genv.find_symbol in H2; simpl in H2. unfold Genv.find_var_info, Genv.find_def in H3; simpl in H3. - rewrite PTree.gsspec in H2. destruct (peq id id1). + rewrite ATree.gsspec in H2. destruct (ATree.elt_eq id id1). + injection H2; clear H2; intro EQ. rewrite EQ, PTree.gss in H3. injection H3; clear H3; intros EQ'; subst v. assert (b = b1). { erewrite Mem.alloc_result by eauto. congruence. } @@ -1751,11 +1751,11 @@ Proof. eapply alloc_global_match; eauto. Qed. -Definition romem_consistent (defmap: PTree.t (globdef fundef unit)) (rm: romem) := +Definition romem_consistent (defmap: ATree.t (globdef fundef unit)) (rm: romem) := forall id ab, - rm!id = Some ab -> + rm$id = Some ab -> exists v, - defmap!id = Some (Gvar v) + defmap$id = Some (Gvar v) /\ v.(gvar_readonly) = true /\ v.(gvar_volatile) = false /\ definitive_initializer v.(gvar_init) = true @@ -1764,18 +1764,18 @@ Definition romem_consistent (defmap: PTree.t (globdef fundef unit)) (rm: romem) Lemma alloc_global_consistent: forall dm rm idg, romem_consistent dm rm -> - romem_consistent (PTree.set (fst idg) (snd idg) dm) (alloc_global rm idg). + romem_consistent (ATree.set (fst idg) (snd idg) dm) (alloc_global rm idg). Proof. intros; red; intros. destruct idg as [id1 [f1 | v1]]; simpl in *. -- rewrite PTree.grspec in H0. destruct (PTree.elt_eq id id1); try discriminate. - rewrite PTree.gso by auto. apply H; auto. +- rewrite ATree.grspec in H0. destruct (ATree.elt_eq id id1); try discriminate. + rewrite ATree.gso by auto. apply H; auto. - destruct (gvar_readonly v1 && negb (gvar_volatile v1) && definitive_initializer (gvar_init v1)) eqn:RO. + InvBooleans. rewrite negb_true_iff in H4. - rewrite PTree.gsspec in *. destruct (peq id id1). + rewrite ATree.gsspec in *. destruct (ATree.elt_eq id id1). * inv H0. exists v1; auto. * apply H; auto. -+ rewrite PTree.grspec in H0. destruct (PTree.elt_eq id id1); try discriminate. - rewrite PTree.gso by auto. apply H; auto. ++ rewrite ATree.grspec in H0. destruct (ATree.elt_eq id id1); try discriminate. + rewrite ATree.gso by auto. apply H; auto. Qed. Lemma romem_for_consistent: @@ -1783,11 +1783,11 @@ Lemma romem_for_consistent: Proof. assert (REC: forall l dm rm, romem_consistent dm rm -> - romem_consistent (fold_left (fun m idg => PTree.set (fst idg) (snd idg) m) l dm) + romem_consistent (fold_left (fun m idg => ATree.set (fst idg) (snd idg) m) l dm) (fold_left alloc_global l rm)). { induction l; intros; simpl; auto. apply IHl. apply alloc_global_consistent; auto. } intros. apply REC. - red; intros. rewrite PTree.gempty in H; discriminate. + red; intros. rewrite ATree.gempty in H; discriminate. Qed. Lemma romem_for_consistent_2: @@ -1822,7 +1822,7 @@ Proof. assert (A: initial_mem_match bc m ge). { apply alloc_globals_match with (m := Mem.empty); auto. - red. unfold Genv.find_symbol; simpl; intros. rewrite PTree.gempty in H1; discriminate. + red. unfold Genv.find_symbol; simpl; intros. rewrite ATree.gempty in H1; discriminate. } assert (B: romem_consistent (prog_defmap prog) (romem_for cunit)) by (apply romem_for_consistent_2; auto). red; intros. diff --git a/backend/ValueDomain.v b/backend/ValueDomain.v index e7e44e29eb..f18d436ceb 100644 --- a/backend/ValueDomain.v +++ b/backend/ValueDomain.v @@ -25,7 +25,7 @@ Inductive block_class : Type := | BCother. Definition block_class_eq: forall (x y: block_class), {x=y} + {x<>y}. -Proof. decide equality. apply peq. Defined. +Proof. decide equality. apply ident_eq. Defined. Record block_classification : Type := BC { bc_img :> block -> block_class; @@ -290,9 +290,9 @@ Qed. Definition pincl (p q: aptr) : bool := match p, q with | Pbot, _ => true - | Gl id1 ofs1, Gl id2 ofs2 => peq id1 id2 && Ptrofs.eq_dec ofs1 ofs2 - | Gl id1 ofs1, Glo id2 => peq id1 id2 - | Glo id1, Glo id2 => peq id1 id2 + | Gl id1 ofs1, Gl id2 ofs2 => ident_eq id1 id2 && Ptrofs.eq_dec ofs1 ofs2 + | Gl id1 ofs1, Glo id2 => ident_eq id1 id2 + | Glo id1, Glo id2 => ident_eq id1 id2 | (Gl _ _ | Glo _ | Glob), Glob => true | (Gl _ _ | Glo _ | Glob | Nonstack), Nonstack => true | Stk ofs1, Stk ofs2 => Ptrofs.eq_dec ofs1 ofs2 @@ -391,14 +391,14 @@ Definition pcmp (c: comparison) (p1 p2: aptr) : abool := match p1, p2 with | Pbot, _ | _, Pbot => Bnone | Gl id1 ofs1, Gl id2 ofs2 => - if peq id1 id2 then Maybe (Ptrofs.cmpu c ofs1 ofs2) + if ident_eq id1 id2 then Maybe (Ptrofs.cmpu c ofs1 ofs2) else cmp_different_blocks c | Gl id1 ofs1, Glo id2 => - if peq id1 id2 then Btop else cmp_different_blocks c + if ident_eq id1 id2 then Btop else cmp_different_blocks c | Glo id1, Gl id2 ofs2 => - if peq id1 id2 then Btop else cmp_different_blocks c + if ident_eq id1 id2 then Btop else cmp_different_blocks c | Glo id1, Glo id2 => - if peq id1 id2 then Btop else cmp_different_blocks c + if ident_eq id1 id2 then Btop else cmp_different_blocks c | Stk ofs1, Stk ofs2 => Maybe (Ptrofs.cmpu c ofs1 ofs2) | (Gl _ _ | Glo _ | Glob | Nonstack), (Stk _ | Stack) => cmp_different_blocks c | (Stk _ | Stack), (Gl _ _ | Glo _ | Glob | Nonstack) => cmp_different_blocks c @@ -435,11 +435,11 @@ Proof. constructor. } unfold pcmp; inv H; inv H0; (apply cmatch_top || (apply DIFF; congruence) || idtac). - - destruct (peq id id0). subst id0. apply SAME. eapply bc_glob; eauto. + - destruct (ident_eq id id0). subst id0. apply SAME. eapply bc_glob; eauto. auto with va. - - destruct (peq id id0); auto with va. - - destruct (peq id id0); auto with va. - - destruct (peq id id0); auto with va. + - destruct (ident_eq id id0); auto with va. + - destruct (ident_eq id id0); auto with va. + - destruct (ident_eq id id0); auto with va. - apply SAME. eapply bc_stack; eauto. Qed. @@ -473,11 +473,11 @@ Proof. constructor. } unfold pcmp; inv H; inv H0; (apply cmatch_top || (apply DIFF; congruence) || idtac). - - destruct (peq id id0). subst id0. apply SAME. eapply bc_glob; eauto. + - destruct (ident_eq id id0). subst id0. apply SAME. eapply bc_glob; eauto. auto with va. - - destruct (peq id id0); auto with va. - - destruct (peq id id0); auto with va. - - destruct (peq id id0); auto with va. + - destruct (ident_eq id id0); auto with va. + - destruct (ident_eq id id0); auto with va. + - destruct (ident_eq id id0); auto with va. - apply SAME. eapply bc_stack; eauto. Qed. @@ -486,7 +486,7 @@ Lemma pcmp_none: Proof. intros. unfold pcmp; destruct p1; try constructor; destruct p2; - try (destruct (peq id id0)); try constructor; try (apply cmp_different_blocks_none). + try (destruct (ident_eq id id0)); try constructor; try (apply cmp_different_blocks_none). Qed. Definition pdisjoint (p1: aptr) (sz1: Z) (p2: aptr) (sz2: Z) : bool := @@ -494,13 +494,13 @@ Definition pdisjoint (p1: aptr) (sz1: Z) (p2: aptr) (sz2: Z) : bool := | Pbot, _ => true | _, Pbot => true | Gl id1 ofs1, Gl id2 ofs2 => - if peq id1 id2 + if ident_eq id1 id2 then zle (Ptrofs.unsigned ofs1 + sz1) (Ptrofs.unsigned ofs2) || zle (Ptrofs.unsigned ofs2 + sz2) (Ptrofs.unsigned ofs1) else true - | Gl id1 ofs1, Glo id2 => negb(peq id1 id2) - | Glo id1, Gl id2 ofs2 => negb(peq id1 id2) - | Glo id1, Glo id2 => negb(peq id1 id2) + | Gl id1 ofs1, Glo id2 => negb(ident_eq id1 id2) + | Glo id1, Gl id2 ofs2 => negb(ident_eq id1 id2) + | Glo id1, Glo id2 => negb(ident_eq id1 id2) | Stk ofs1, Stk ofs2 => zle (Ptrofs.unsigned ofs1 + sz1) (Ptrofs.unsigned ofs2) || zle (Ptrofs.unsigned ofs2 + sz2) (Ptrofs.unsigned ofs1) @@ -516,11 +516,11 @@ Lemma pdisjoint_sound: b1 <> b2 \/ Ptrofs.unsigned ofs1 + sz1 <= Ptrofs.unsigned ofs2 \/ Ptrofs.unsigned ofs2 + sz2 <= Ptrofs.unsigned ofs1. Proof. intros. inv H0; inv H1; simpl in H; try discriminate; try (left; congruence). -- destruct (peq id id0). subst id0. destruct (orb_true_elim _ _ H); InvBooleans; auto. +- destruct (ident_eq id id0). subst id0. destruct (orb_true_elim _ _ H); InvBooleans; auto. left; congruence. -- destruct (peq id id0); try discriminate. left; congruence. -- destruct (peq id id0); try discriminate. left; congruence. -- destruct (peq id id0); try discriminate. left; congruence. +- destruct (ident_eq id id0); try discriminate. left; congruence. +- destruct (ident_eq id id0); try discriminate. left; congruence. +- destruct (ident_eq id id0); try discriminate. left; congruence. - destruct (orb_true_elim _ _ H); InvBooleans; auto. Qed. @@ -3616,12 +3616,12 @@ Qed. (** * Abstracting read-only global variables *) -Definition romem := PTree.t ablock. +Definition romem := ATree.t ablock. Definition romatch (m: mem) (rm: romem) : Prop := forall b id ab, bc b = BCglob id -> - rm!id = Some ab -> + rm$id = Some ab -> pge Glob ab.(ab_summary) /\ bmatch m b ab /\ forall ofs, ~Mem.perm m b ofs Max Writable. @@ -3697,7 +3697,7 @@ Qed. Record amem : Type := AMem { am_stack: ablock; - am_glob: PTree.t ablock; + am_glob: ATree.t ablock; am_nonstack: aptr; am_top: aptr }. @@ -3708,7 +3708,7 @@ Record mmatch (m: mem) (am: amem) : Prop := mk_mem_match { bmatch m b am.(am_stack); mmatch_glob: forall id ab b, bc b = BCglob id -> - am.(am_glob)!id = Some ab -> + am.(am_glob)$id = Some ab -> bmatch m b ab; mmatch_nonstack: forall b, bc b <> BCstack -> bc b <> BCinvalid -> @@ -3722,7 +3722,7 @@ Record mmatch (m: mem) (am: amem) : Prop := mk_mem_match { Definition minit (p: aptr) := {| am_stack := ablock_init p; - am_glob := PTree.empty _; + am_glob := ATree.empty _; am_nonstack := p; am_top := p |}. @@ -3733,19 +3733,19 @@ Definition load (chunk: memory_chunk) (rm: romem) (m: amem) (p: aptr) : aval := match p with | Pbot => if va_strict tt then Vbot else Vtop | Gl id ofs => - match rm!id with + match rm$id with | Some ab => ablock_load chunk ab (Ptrofs.unsigned ofs) | None => - match m.(am_glob)!id with + match m.(am_glob)$id with | Some ab => ablock_load chunk ab (Ptrofs.unsigned ofs) | None => vnormalize chunk (Ifptr m.(am_nonstack)) end end | Glo id => - match rm!id with + match rm$id with | Some ab => ablock_load_anywhere chunk ab | None => - match m.(am_glob)!id with + match m.(am_glob)$id with | Some ab => ablock_load_anywhere chunk ab | None => vnormalize chunk (Ifptr m.(am_nonstack)) end @@ -3769,12 +3769,12 @@ Definition store (chunk: memory_chunk) (m: amem) (p: aptr) (av: aval) : amem := am_glob := match p with | Gl id ofs => - let ab := match m.(am_glob)!id with Some ab => ab | None => ablock_init m.(am_nonstack) end in - PTree.set id (ablock_store chunk ab (Ptrofs.unsigned ofs) av) m.(am_glob) + let ab := match m.(am_glob)$id with Some ab => ab | None => ablock_init m.(am_nonstack) end in + ATree.set id (ablock_store chunk ab (Ptrofs.unsigned ofs) av) m.(am_glob) | Glo id => - let ab := match m.(am_glob)!id with Some ab => ab | None => ablock_init m.(am_nonstack) end in - PTree.set id (ablock_store_anywhere chunk ab av) m.(am_glob) - | Glob | Nonstack | Ptop => PTree.empty _ + let ab := match m.(am_glob)$id with Some ab => ab | None => ablock_init m.(am_nonstack) end in + ATree.set id (ablock_store_anywhere chunk ab av) m.(am_glob) + | Glob | Nonstack | Ptop => ATree.empty _ | _ => m.(am_glob) end; am_nonstack := @@ -3792,10 +3792,10 @@ Definition loadbytes (m: amem) (rm: romem) (p: aptr) : aptr := match p with | Pbot => if va_strict tt then Pbot else Ptop | Gl id _ | Glo id => - match rm!id with + match rm$id with | Some ab => ablock_loadbytes ab | None => - match m.(am_glob)!id with + match m.(am_glob)$id with | Some ab => ablock_loadbytes ab | None => m.(am_nonstack) end @@ -3815,12 +3815,12 @@ Definition storebytes (m: amem) (dst: aptr) (sz: Z) (p: aptr) : amem := am_glob := match dst with | Gl id ofs => - let ab := match m.(am_glob)!id with Some ab => ab | None => ablock_init m.(am_nonstack) end in - PTree.set id (ablock_storebytes ab p (Ptrofs.unsigned ofs) sz) m.(am_glob) + let ab := match m.(am_glob)$id with Some ab => ab | None => ablock_init m.(am_nonstack) end in + ATree.set id (ablock_storebytes ab p (Ptrofs.unsigned ofs) sz) m.(am_glob) | Glo id => - let ab := match m.(am_glob)!id with Some ab => ab | None => ablock_init m.(am_nonstack) end in - PTree.set id (ablock_storebytes_anywhere ab p) m.(am_glob) - | Glob | Nonstack | Ptop => PTree.empty _ + let ab := match m.(am_glob)$id with Some ab => ab | None => ablock_init m.(am_nonstack) end in + ATree.set id (ablock_storebytes_anywhere ab p) m.(am_glob) + | Glob | Nonstack | Ptop => ATree.empty _ | _ => m.(am_glob) end; am_nonstack := @@ -3841,15 +3841,15 @@ Theorem load_sound: Proof. intros. unfold load. inv H2. - (* Gl id ofs *) - destruct (rm!id) as [ab|] eqn:RM. + destruct (rm$id) as [ab|] eqn:RM. eapply ablock_load_sound; eauto. eapply H0; eauto. - destruct (am_glob am)!id as [ab|] eqn:AM. + destruct (am_glob am)$id as [ab|] eqn:AM. eapply ablock_load_sound; eauto. eapply mmatch_glob; eauto. eapply vnormalize_cast; eauto. eapply mmatch_nonstack; eauto; congruence. - (* Glo id *) - destruct (rm!id) as [ab|] eqn:RM. + destruct (rm$id) as [ab|] eqn:RM. eapply ablock_load_anywhere_sound; eauto. eapply H0; eauto. - destruct (am_glob am)!id as [ab|] eqn:AM. + destruct (am_glob am)$id as [ab|] eqn:AM. eapply ablock_load_anywhere_sound; eauto. eapply mmatch_glob; eauto. eapply vnormalize_cast; eauto. eapply mmatch_nonstack; eauto; congruence. - (* Glob *) @@ -3899,32 +3899,32 @@ Proof. - (* Globals *) rename b0 into b'. - assert (DFL: bc b <> BCglob id -> (am_glob am)!id = Some ab -> + assert (DFL: bc b <> BCglob id -> (am_glob am)$id = Some ab -> bmatch m' b' ab). { intros. apply bmatch_inv with m. eapply mmatch_glob; eauto. intros. eapply Mem.loadbytes_store_other; eauto. left; congruence. } inv PM. - + rewrite PTree.gsspec in H0. destruct (peq id id0). + + rewrite ATree.gsspec in H0. destruct (ATree.elt_eq id id0). subst id0; inv H0. assert (b' = b) by (eapply bc_glob; eauto). subst b'. eapply ablock_store_sound; eauto. - destruct (am_glob am)!id as [ab0|] eqn:GL. + destruct (am_glob am)$id as [ab0|] eqn:GL. eapply mmatch_glob; eauto. apply ablock_init_sound. eapply mmatch_nonstack; eauto; congruence. eapply DFL; eauto. congruence. - + rewrite PTree.gsspec in H0. destruct (peq id id0). + + rewrite ATree.gsspec in H0. destruct (ATree.elt_eq id id0). subst id0; inv H0. assert (b' = b) by (eapply bc_glob; eauto). subst b'. eapply ablock_store_anywhere_sound; eauto. - destruct (am_glob am)!id as [ab0|] eqn:GL. + destruct (am_glob am)$id as [ab0|] eqn:GL. eapply mmatch_glob; eauto. apply ablock_init_sound. eapply mmatch_nonstack; eauto; congruence. eapply DFL; eauto. congruence. - + rewrite PTree.gempty in H0; congruence. + + rewrite ATree.gempty in H0; congruence. + eapply DFL; eauto. congruence. + eapply DFL; eauto. congruence. - + rewrite PTree.gempty in H0; congruence. - + rewrite PTree.gempty in H0; congruence. + + rewrite ATree.gempty in H0; congruence. + + rewrite ATree.gempty in H0; congruence. - (* Nonstack *) assert (DFL: smatch m' b0 (vplub av (am_nonstack am))). @@ -3964,15 +3964,15 @@ Theorem loadbytes_sound: Proof. intros. unfold loadbytes; inv H2. - (* Gl id ofs *) - destruct (rm!id) as [ab|] eqn:RM. + destruct (rm$id) as [ab|] eqn:RM. exploit H0; eauto. intros (A & B & C). eapply ablock_loadbytes_sound; eauto. - destruct (am_glob am)!id as [ab|] eqn:GL. + destruct (am_glob am)$id as [ab|] eqn:GL. eapply ablock_loadbytes_sound; eauto. eapply mmatch_glob; eauto. eapply smatch_loadbytes; eauto. eapply mmatch_nonstack; eauto with va. - (* Glo id *) - destruct (rm!id) as [ab|] eqn:RM. + destruct (rm$id) as [ab|] eqn:RM. exploit H0; eauto. intros (A & B & C). eapply ablock_loadbytes_sound; eauto. - destruct (am_glob am)!id as [ab|] eqn:GL. + destruct (am_glob am)$id as [ab|] eqn:GL. eapply ablock_loadbytes_sound; eauto. eapply mmatch_glob; eauto. eapply smatch_loadbytes; eauto. eapply mmatch_nonstack; eauto with va. - (* Glob *) @@ -4011,32 +4011,32 @@ Proof. - (* Globals *) rename b0 into b'. - assert (DFL: bc b <> BCglob id -> (am_glob am)!id = Some ab -> + assert (DFL: bc b <> BCglob id -> (am_glob am)$id = Some ab -> bmatch m' b' ab). { intros. apply bmatch_inv with m. eapply mmatch_glob; eauto. intros. eapply Mem.loadbytes_storebytes_other; eauto. left; congruence. } inv PM. - + rewrite PTree.gsspec in H0. destruct (peq id id0). + + rewrite ATree.gsspec in H0. destruct (ATree.elt_eq id id0). subst id0; inv H0. assert (b' = b) by (eapply bc_glob; eauto). subst b'. eapply ablock_storebytes_sound; eauto. - destruct (am_glob am)!id as [ab0|] eqn:GL. + destruct (am_glob am)$id as [ab0|] eqn:GL. eapply mmatch_glob; eauto. apply ablock_init_sound. eapply mmatch_nonstack; eauto; congruence. eapply DFL; eauto. congruence. - + rewrite PTree.gsspec in H0. destruct (peq id id0). + + rewrite ATree.gsspec in H0. destruct (ATree.elt_eq id id0). subst id0; inv H0. assert (b' = b) by (eapply bc_glob; eauto). subst b'. eapply ablock_storebytes_anywhere_sound; eauto. - destruct (am_glob am)!id as [ab0|] eqn:GL. + destruct (am_glob am)$id as [ab0|] eqn:GL. eapply mmatch_glob; eauto. apply ablock_init_sound. eapply mmatch_nonstack; eauto; congruence. eapply DFL; eauto. congruence. - + rewrite PTree.gempty in H0; congruence. + + rewrite ATree.gempty in H0; congruence. + eapply DFL; eauto. congruence. + eapply DFL; eauto. congruence. - + rewrite PTree.gempty in H0; congruence. - + rewrite PTree.gempty in H0; congruence. + + rewrite ATree.gempty in H0; congruence. + + rewrite ATree.gempty in H0; congruence. - (* Nonstack *) assert (DFL: smatch m' b0 (plub q (am_nonstack am))). @@ -4086,7 +4086,7 @@ Proof. intros. constructor; simpl; intros. - apply ablock_init_sound. apply smatch_ge with (ab_summary (am_stack am)). eapply mmatch_stack; eauto. constructor. -- rewrite PTree.gempty in H1; discriminate. +- rewrite ATree.gempty in H1; discriminate. - eapply smatch_ge. eapply mmatch_nonstack; eauto. constructor. - eapply smatch_ge. eapply mmatch_top; eauto. constructor. - eapply mmatch_below; eauto. @@ -4098,21 +4098,21 @@ Definition mbeq (m1 m2: amem) : bool := eq_aptr m1.(am_top) m2.(am_top) && eq_aptr m1.(am_nonstack) m2.(am_nonstack) && bbeq m1.(am_stack) m2.(am_stack) - && PTree.beq bbeq m1.(am_glob) m2.(am_glob). + && ATree.beq bbeq m1.(am_glob) m2.(am_glob). Lemma mbeq_sound: forall m1 m2, mbeq m1 m2 = true -> forall m, mmatch m m1 <-> mmatch m m2. Proof. - unfold mbeq; intros. InvBooleans. rewrite PTree.beq_correct in H1. + unfold mbeq; intros. InvBooleans. rewrite ATree.beq_correct in H1. split; intros M; inv M; constructor; intros. - erewrite <- bbeq_sound; eauto. -- specialize (H1 id). rewrite H4 in H1. destruct (am_glob m1)!id eqn:G; try contradiction. +- specialize (H1 id). rewrite H4 in H1. destruct (am_glob m1)$id eqn:G; try contradiction. erewrite <- bbeq_sound; eauto. - rewrite <- H; eauto. - rewrite <- H0; eauto. - auto. - erewrite bbeq_sound; eauto. -- specialize (H1 id). rewrite H4 in H1. destruct (am_glob m2)!id eqn:G; try contradiction. +- specialize (H1 id). rewrite H4 in H1. destruct (am_glob m2)$id eqn:G; try contradiction. erewrite bbeq_sound; eauto. - rewrite H; eauto. - rewrite H0; eauto. @@ -4129,7 +4129,7 @@ Definition combine_ablock (ob1 ob2: option ablock) : option ablock := Definition mlub (m1 m2: amem) : amem := {| am_stack := blub m1.(am_stack) m2.(am_stack); - am_glob := PTree.combine combine_ablock m1.(am_glob) m2.(am_glob); + am_glob := ATree.combine combine_ablock m1.(am_glob) m2.(am_glob); am_nonstack := plub m1.(am_nonstack) m2.(am_nonstack); am_top := plub m1.(am_top) m2.(am_top) |}. @@ -4138,9 +4138,9 @@ Lemma mmatch_lub_l: Proof. intros. inv H. constructor; simpl; intros. - apply bmatch_lub_l; auto. -- rewrite PTree.gcombine in H0 by auto. unfold combine_ablock in H0. - destruct (am_glob x)!id as [b1|] eqn:G1; - destruct (am_glob y)!id as [b2|] eqn:G2; +- rewrite ATree.gcombine in H0 by auto. unfold combine_ablock in H0. + destruct (am_glob x)$id as [b1|] eqn:G1; + destruct (am_glob y)$id as [b2|] eqn:G2; inv H0. apply bmatch_lub_l; eauto. - apply smatch_lub_l; auto. @@ -4153,9 +4153,9 @@ Lemma mmatch_lub_r: Proof. intros. inv H. constructor; simpl; intros. - apply bmatch_lub_r; auto. -- rewrite PTree.gcombine in H0 by auto. unfold combine_ablock in H0. - destruct (am_glob x)!id as [b1|] eqn:G1; - destruct (am_glob y)!id as [b2|] eqn:G2; +- rewrite ATree.gcombine in H0 by auto. unfold combine_ablock in H0. + destruct (am_glob x)$id as [b1|] eqn:G1; + destruct (am_glob y)$id as [b2|] eqn:G2; inv H0. apply bmatch_lub_r; eauto. - apply smatch_lub_r; auto. @@ -4364,7 +4364,7 @@ Proof. } constructor; simpl; intros. - apply ablock_init_sound. apply SM. congruence. - - rewrite PTree.gempty in H1; discriminate. + - rewrite ATree.gempty in H1; discriminate. - apply SM; auto. - apply SM; auto. - red; intros. eapply Mem.valid_block_inject_1. eapply inj_of_bc_valid; eauto. eauto. diff --git a/cfrontend/C2C.ml b/cfrontend/C2C.ml index b81345991b..6480a337e8 100644 --- a/cfrontend/C2C.ml +++ b/cfrontend/C2C.ml @@ -107,7 +107,7 @@ let atom_location a = (** The current environment of composite definitions *) -let comp_env : composite_env ref = ref Maps.PTree.empty +let comp_env : composite_env ref = ref Symbols.ATree.empty (** Hooks -- overriden in machine-dependent CPragmas module *) @@ -1410,7 +1410,7 @@ let convertProgram p = comp_env := ce; let gl1 = convertGlobdecls env [] p in let gl2 = globals_for_strings gl1 in - comp_env := Maps.PTree.empty; + comp_env := Symbols.ATree.empty; let p' = { prog_defs = gl2; prog_public = public_globals gl2; diff --git a/cfrontend/Cexec.v b/cfrontend/Cexec.v index 823d2542d9..9f7e5205e3 100644 --- a/cfrontend/Cexec.v +++ b/cfrontend/Cexec.v @@ -154,7 +154,7 @@ Proof. - auto. - auto. - auto. -- rewrite (Genv.find_invert_symbol _ _ H0). simpl in H; rewrite H. +- rewrite (Genv.find_invert_symbol H0). simpl in H; rewrite H. rewrite dec_eq_true. auto. Qed. @@ -238,7 +238,7 @@ Lemma do_volatile_load_complete: do_volatile_load w chunk m b ofs = Some(w', t, v). Proof. unfold do_volatile_load; intros. inv H; simpl in *. - rewrite H1. rewrite (Genv.find_invert_symbol _ _ H2). inv H0. inv H8. inv H6. rewrite H9. + rewrite H1. rewrite (Genv.find_invert_symbol H2). inv H0. inv H8. inv H6. rewrite H9. rewrite (val_of_eventval_complete _ _ _ H3). auto. rewrite H1. rewrite H2. inv H0. auto. Qed. @@ -261,7 +261,7 @@ Lemma do_volatile_store_complete: do_volatile_store w chunk m b ofs v = Some(w', t, m'). Proof. unfold do_volatile_store; intros. inv H; simpl in *. - rewrite H1. rewrite (Genv.find_invert_symbol _ _ H2). + rewrite H1. rewrite (Genv.find_invert_symbol H2). rewrite (eventval_of_val_complete _ _ _ H3). inv H0. inv H8. inv H6. rewrite H9. auto. rewrite H1. rewrite H2. inv H0. auto. @@ -682,7 +682,7 @@ Fixpoint step_expr (k: kind) (a: expr) (m: mem): reducts expr := | LV, Eloc b ofs ty => nil | LV, Evar x ty => - match e!x with + match e$x with | Some(b, ty') => check type_eq ty ty'; topred (Lred "red_var_local" (Eloc b Ptrofs.zero ty) m) @@ -704,13 +704,13 @@ Fixpoint step_expr (k: kind) (a: expr) (m: mem): reducts expr := | Some(Vptr b ofs, ty') => match ty' with | Tstruct id _ => - do co <- ge.(genv_cenv)!id; + do co <- ge.(genv_cenv)$id; match field_offset ge f (co_members co) with | Error _ => stuck | OK delta => topred (Lred "red_field_struct" (Eloc b (Ptrofs.add ofs (Ptrofs.repr delta)) ty) m) end | Tunion id _ => - do co <- ge.(genv_cenv)!id; + do co <- ge.(genv_cenv)$id; topred (Lred "red_field_union" (Eloc b ofs ty) m) | _ => stuck end @@ -925,15 +925,15 @@ Definition invert_expr_prop (a: expr) (m: mem) : Prop := | Eloc b ofs ty => False | Evar x ty => exists b, - e!x = Some(b, ty) - \/ (e!x = None /\ Genv.find_symbol ge x = Some b) + e$x = Some(b, ty) + \/ (e$x = None /\ Genv.find_symbol ge x = Some b) | Ederef (Eval v ty1) ty => exists b, exists ofs, v = Vptr b ofs | Efield (Eval v ty1) f ty => exists b, exists ofs, v = Vptr b ofs /\ match ty1 with - | Tstruct id _ => exists co delta, ge.(genv_cenv)!id = Some co /\ field_offset ge f (co_members co) = OK delta - | Tunion id _ => exists co, ge.(genv_cenv)!id = Some co + | Tstruct id _ => exists co delta, ge.(genv_cenv)$id = Some co /\ field_offset ge f (co_members co) = OK delta + | Tunion id _ => exists co, ge.(genv_cenv)$id = Some co | _ => False end | Eval v ty => False @@ -1318,7 +1318,7 @@ Proof with (try (apply not_invert_ok; simpl; intro; myinv; intuition congruence; (* Eval *) split; intros. tauto. simpl; congruence. (* Evar *) - destruct (e!x) as [[b ty']|] eqn:?. + destruct (e$x) as [[b ty']|] eqn:?. destruct (type_eq ty ty')... subst. apply topred_ok; auto. apply red_var_local; auto. destruct (Genv.find_symbol ge x) as [b|] eqn:?... @@ -1329,11 +1329,11 @@ Proof with (try (apply not_invert_ok; simpl; intro; myinv; intuition congruence; destruct v... destruct ty'... (* top struct *) - destruct (ge.(genv_cenv)!i0) as [co|] eqn:?... + destruct (ge.(genv_cenv)$i0) as [co|] eqn:?... destruct (field_offset ge f (co_members co)) as [delta|] eqn:?... apply topred_ok; auto. eapply red_field_struct; eauto. (* top union *) - destruct (ge.(genv_cenv)!i0) as [co|] eqn:?... + destruct (ge.(genv_cenv)$i0) as [co|] eqn:?... apply topred_ok; auto. eapply red_field_union; eauto. (* in depth *) eapply incontext_ok; eauto. @@ -1846,7 +1846,7 @@ Fixpoint do_alloc_variables (e: env) (m: mem) (l: list (ident * type)) {struct l | nil => (e,m) | (id, ty) :: l' => let (m1,b1) := Mem.alloc m 0 (sizeof ge ty) in - do_alloc_variables (PTree.set id (b1, ty) e) m1 l' + do_alloc_variables (ATree.set id (b1, ty) e) m1 l' end. Lemma do_alloc_variables_sound: @@ -1872,7 +1872,7 @@ Function sem_bind_parameters (w: world) (e: env) (m: mem) (l: list (ident * type match l, lv with | nil, nil => Some m | (id, ty) :: params, v1::lv => - match PTree.get id e with + match ATree.get id e with | Some (b, ty') => check (type_eq ty ty'); do w', t, m1 <- do_assign_loc w ty m b Ptrofs.zero v1; diff --git a/cfrontend/Clight.v b/cfrontend/Clight.v index 7a4c49a2f3..55f1877075 100644 --- a/cfrontend/Clight.v +++ b/cfrontend/Clight.v @@ -188,13 +188,13 @@ Definition globalenv (p: program) := types. The current value of the variable is stored in the associated memory block. *) -Definition env := PTree.t (block * type). (* map variable -> location & type *) +Definition env := ATree.t (block * type). (* map variable -> location & type *) -Definition empty_env: env := (PTree.empty (block * type)). +Definition empty_env: env := (ATree.empty (block * type)). (** The temporary environment maps local temporaries to values. *) -Definition temp_env := PTree.t val. +Definition temp_env := ATree.t val. (** [deref_loc ty m b ofs v] computes the value of a datum of type [ty] residing in memory [m] at block [b], offset [ofs]. @@ -257,7 +257,7 @@ Inductive alloc_variables: env -> mem -> | alloc_variables_cons: forall e m id ty vars m1 b1 m2 e2, Mem.alloc m 0 (sizeof ge ty) = (m1, b1) -> - alloc_variables (PTree.set id (b1, ty) e) m1 vars e2 m2 -> + alloc_variables (ATree.set id (b1, ty) e) m1 vars e2 m2 -> alloc_variables e m ((id, ty) :: vars) e2 m2. (** Initialization of local variables that are parameters to a function. @@ -273,7 +273,7 @@ Inductive bind_parameters (e: env): bind_parameters e m nil nil m | bind_parameters_cons: forall m id ty params v1 vl b m1 m2, - PTree.get id e = Some(b, ty) -> + ATree.get id e = Some(b, ty) -> assign_loc ge ty m b Ptrofs.zero v1 m1 -> bind_parameters e m1 params vl m2 -> bind_parameters e m ((id, ty) :: params) (v1 :: vl) m2. @@ -282,8 +282,8 @@ Inductive bind_parameters (e: env): Fixpoint create_undef_temps (temps: list (ident * type)) : temp_env := match temps with - | nil => PTree.empty val - | (id, t) :: temps' => PTree.set id Vundef (create_undef_temps temps') + | nil => ATree.empty val + | (id, t) :: temps' => ATree.set id Vundef (create_undef_temps temps') end. (** Initialization of temporary variables that are parameters to a function. *) @@ -292,7 +292,7 @@ Fixpoint bind_parameter_temps (formals: list (ident * type)) (args: list val) (le: temp_env) : option temp_env := match formals, args with | nil, nil => Some le - | (id, t) :: xl, v :: vl => bind_parameter_temps xl vl (PTree.set id v le) + | (id, t) :: xl, v :: vl => bind_parameter_temps xl vl (ATree.set id v le) | _, _ => None end. @@ -302,14 +302,14 @@ Definition block_of_binding (id_b_ty: ident * (block * type)) := match id_b_ty with (id, (b, ty)) => (b, 0, sizeof ge ty) end. Definition blocks_of_env (e: env) : list (block * Z * Z) := - List.map block_of_binding (PTree.elements e). + List.map block_of_binding (ATree.elements e). (** Optional assignment to a temporary *) Definition set_opttemp (optid: option ident) (v: val) (le: temp_env) := match optid with | None => le - | Some id => PTree.set id v le + | Some id => ATree.set id v le end. (** Selection of the appropriate case of a [switch], given the value [n] @@ -365,7 +365,7 @@ Inductive eval_expr: expr -> val -> Prop := | eval_Econst_long: forall i ty, eval_expr (Econst_long i ty) (Vlong i) | eval_Etempvar: forall id ty v, - le!id = Some v -> + le$id = Some v -> eval_expr (Etempvar id ty) v | eval_Eaddrof: forall a ty loc ofs, eval_lvalue a loc ofs -> @@ -398,10 +398,10 @@ Inductive eval_expr: expr -> val -> Prop := with eval_lvalue: expr -> block -> ptrofs -> Prop := | eval_Evar_local: forall id l ty, - e!id = Some(l, ty) -> + e$id = Some(l, ty) -> eval_lvalue (Evar id ty) l Ptrofs.zero | eval_Evar_global: forall id l ty, - e!id = None -> + e$id = None -> Genv.find_symbol ge id = Some l -> eval_lvalue (Evar id ty) l Ptrofs.zero | eval_Ederef: forall a ty l ofs, @@ -410,13 +410,13 @@ with eval_lvalue: expr -> block -> ptrofs -> Prop := | eval_Efield_struct: forall a i ty l ofs id co att delta, eval_expr a (Vptr l ofs) -> typeof a = Tstruct id att -> - ge.(genv_cenv)!id = Some co -> + ge.(genv_cenv)$id = Some co -> field_offset ge i (co_members co) = OK delta -> eval_lvalue (Efield a i ty) l (Ptrofs.add ofs (Ptrofs.repr delta)) | eval_Efield_union: forall a i ty l ofs id co att, eval_expr a (Vptr l ofs) -> typeof a = Tunion id att -> - ge.(genv_cenv)!id = Some co -> + ge.(genv_cenv)$id = Some co -> eval_lvalue (Efield a i ty) l ofs. Scheme eval_expr_ind2 := Minimality for eval_expr Sort Prop @@ -557,7 +557,7 @@ Inductive step: state -> trace -> state -> Prop := | step_set: forall f id a k e le m v, eval_expr e le m a v -> step (State f (Sset id a) k e le m) - E0 (State f Sskip k e (PTree.set id v le) m) + E0 (State f Sskip k e (ATree.set id v le) m) | step_call: forall f optid a al k e le m tyargs tyres cconv vf vargs fd, classify_fun (typeof a) = fun_case_f tyargs tyres cconv -> diff --git a/cfrontend/ClightBigstep.v b/cfrontend/ClightBigstep.v index 92457586eb..2d8869c646 100644 --- a/cfrontend/ClightBigstep.v +++ b/cfrontend/ClightBigstep.v @@ -88,7 +88,7 @@ Inductive exec_stmt: env -> temp_env -> mem -> statement -> trace -> temp_env -> | exec_Sset: forall e le m id a v, eval_expr ge e le m a v -> exec_stmt e le m (Sset id a) - E0 (PTree.set id v le) m Out_normal + E0 (ATree.set id v le) m Out_normal | exec_Scall: forall e le m optid a al tyargs tyres cconv vf vargs f t m' vres, classify_fun (typeof a) = fun_case_f tyargs tyres cconv -> eval_expr ge e le m a vf -> diff --git a/cfrontend/Cminorgen.v b/cfrontend/Cminorgen.v index 45c21f9677..29bf470c3f 100644 --- a/cfrontend/Cminorgen.v +++ b/cfrontend/Cminorgen.v @@ -40,13 +40,13 @@ Local Open Scope error_monad_scope. (** To every Csharpminor variable, the compiler associates its offset in the Cminor stack data block. *) -Definition compilenv := PTree.t Z. +Definition compilenv := ATree.t Z. (** Generation of a Cminor expression for taking the address of a Csharpminor variable. *) Definition var_addr (cenv: compilenv) (id: ident): expr := - match PTree.get id cenv with + match ATree.get id cenv with | Some ofs => Econst (Oaddrstack (Ptrofs.repr ofs)) | None => Econst (Oaddrsymbol id Ptrofs.zero) end. @@ -226,7 +226,7 @@ Definition assign_variable let (id, sz) := id_sz in let (cenv, stacksize) := cenv_stacksize in let ofs := align stacksize (block_alignment sz) in - (PTree.set id ofs cenv, ofs + Z.max 0 sz). + (ATree.set id ofs cenv, ofs + Z.max 0 sz). Definition assign_variables (cenv_stacksize: compilenv * Z) (vars: list (ident * Z)) : compilenv * Z := List.fold_left assign_variable vars cenv_stacksize. @@ -248,7 +248,7 @@ End VarOrder. Module VarSort := Mergesort.Sort(VarOrder). Definition build_compilenv (f: Csharpminor.function) : compilenv * Z := - assign_variables (PTree.empty Z, 0) (VarSort.sort (Csharpminor.fn_vars f)). + assign_variables (ATree.empty Z, 0) (VarSort.sort (Csharpminor.fn_vars f)). (** * Translation of functions *) diff --git a/cfrontend/Cminorgenproof.v b/cfrontend/Cminorgenproof.v index ffafc5d242..520bafa6f0 100644 --- a/cfrontend/Cminorgenproof.v +++ b/cfrontend/Cminorgenproof.v @@ -164,7 +164,7 @@ Qed. (** ** Matching between Cshaprminor's temporaries and Cminor's variables *) Definition match_temps (f: meminj) (le: Csharpminor.temp_env) (te: env) : Prop := - forall id v, le!id = Some v -> exists v', te!(id) = Some v' /\ Val.inject f v v'. + forall id v, le$id = Some v -> exists v', te$(id) = Some v' /\ Val.inject f v v'. Lemma match_temps_invariant: forall f f' le te, @@ -179,9 +179,9 @@ Lemma match_temps_assign: forall f le te id v tv, match_temps f le te -> Val.inject f v tv -> - match_temps f (PTree.set id v le) (PTree.set id tv te). + match_temps f (ATree.set id v le) (ATree.set id tv te). Proof. - intros; red; intros. rewrite PTree.gsspec in *. destruct (peq id0 id). + intros; red; intros. rewrite ATree.gsspec in *. destruct (ATree.elt_eq id0 id). inv H1. exists tv; auto. eauto. Qed. @@ -206,7 +206,7 @@ Record match_env (f: meminj) (cenv: compilenv) (** C#minor local variables match sub-blocks of the Cminor stack data block. *) me_vars: - forall id, match_var f sp (e!id) (cenv!id); + forall id, match_var f sp (e$id) (cenv$id); (** [lo, hi] is a proper interval. *) me_low_high: @@ -215,14 +215,14 @@ Record match_env (f: meminj) (cenv: compilenv) (** Every block appearing in the C#minor environment [e] must be in the range [lo, hi]. *) me_bounded: - forall id b sz, PTree.get id e = Some(b, sz) -> Ple lo b /\ Plt b hi; + forall id b sz, ATree.get id e = Some(b, sz) -> Ple lo b /\ Plt b hi; (** All blocks mapped to sub-blocks of the Cminor stack data must be images of variables from the C#minor environment [e] *) me_inv: forall b delta, f b = Some(sp, delta) -> - exists id, exists sz, PTree.get id e = Some(b, sz); + exists id, exists sz, ATree.get id e = Some(b, sz); (** All C#minor blocks below [lo] (i.e. allocated before the blocks referenced from [e]) must map to blocks that are below [sp] @@ -294,39 +294,39 @@ Qed. Lemma match_env_alloc: forall f1 id cenv e sp lo m1 sz m2 b ofs f2, - match_env f1 (PTree.remove id cenv) e sp lo (Mem.nextblock m1) -> + match_env f1 (ATree.remove id cenv) e sp lo (Mem.nextblock m1) -> Mem.alloc m1 0 sz = (m2, b) -> - cenv!id = Some ofs -> + cenv$id = Some ofs -> inject_incr f1 f2 -> f2 b = Some(sp, ofs) -> (forall b', b' <> b -> f2 b' = f1 b') -> - e!id = None -> - match_env f2 cenv (PTree.set id (b, sz) e) sp lo (Mem.nextblock m2). + e$id = None -> + match_env f2 cenv (ATree.set id (b, sz) e) sp lo (Mem.nextblock m2). Proof. intros until f2; intros ME ALLOC CENV INCR SAME OTHER ENV. exploit Mem.nextblock_alloc; eauto. intros NEXTBLOCK. exploit Mem.alloc_result; eauto. intros RES. inv ME; constructor. (* vars *) - intros. rewrite PTree.gsspec. destruct (peq id0 id). + intros. rewrite ATree.gsspec. destruct (ATree.elt_eq id0 id). (* the new var *) subst id0. rewrite CENV. constructor. econstructor. eauto. rewrite Ptrofs.add_commut; rewrite Ptrofs.add_zero; auto. (* old vars *) - generalize (me_vars0 id0). rewrite PTree.gro; auto. intros M; inv M. + generalize (me_vars0 id0). rewrite ATree.gro; auto. intros M; inv M. constructor; eauto. constructor. (* low-high *) rewrite NEXTBLOCK; xomega. (* bounded *) - intros. rewrite PTree.gsspec in H. destruct (peq id0 id). + intros. rewrite ATree.gsspec in H. destruct (ATree.elt_eq id0 id). inv H. rewrite NEXTBLOCK; xomega. exploit me_bounded0; eauto. rewrite NEXTBLOCK; xomega. (* inv *) intros. destruct (eq_block b (Mem.nextblock m1)). - subst b. rewrite SAME in H; inv H. exists id; exists sz. apply PTree.gss. + subst b. rewrite SAME in H; inv H. exists id; exists sz. apply ATree.gss. rewrite OTHER in H; auto. exploit me_inv0; eauto. - intros [id1 [sz1 EQ]]. exists id1; exists sz1. rewrite PTree.gso; auto. congruence. + intros [id1 [sz1 EQ]]. exists id1; exists sz1. rewrite ATree.gso; auto. congruence. (* incr *) intros. rewrite OTHER in H. eauto. unfold block in *; xomega. Qed. @@ -335,13 +335,13 @@ Qed. Definition match_bounds (e: Csharpminor.env) (m: mem) : Prop := forall id b sz ofs p, - PTree.get id e = Some(b, sz) -> Mem.perm m b ofs Max p -> 0 <= ofs < sz. + ATree.get id e = Some(b, sz) -> Mem.perm m b ofs Max p -> 0 <= ofs < sz. Lemma match_bounds_invariant: forall e m1 m2, match_bounds e m1 -> (forall id b sz ofs p, - PTree.get id e = Some(b, sz) -> Mem.perm m2 b ofs Max p -> Mem.perm m1 b ofs Max p) -> + ATree.get id e = Some(b, sz) -> Mem.perm m2 b ofs Max p -> Mem.perm m1 b ofs Max p) -> match_bounds e m2. Proof. intros; red; intros. eapply H; eauto. @@ -354,7 +354,7 @@ Qed. Inductive is_reachable_from_env (f: meminj) (e: Csharpminor.env) (sp: block) (ofs: Z) : Prop := | is_reachable_intro: forall id b sz delta, - e!id = Some(b, sz) -> + e$id = Some(b, sz) -> f b = Some(sp, delta) -> delta <= ofs < delta + sz -> is_reachable_from_env f e sp ofs. @@ -395,7 +395,7 @@ Proof. else false end end). - destruct (List.existsb pred (PTree.elements e)) eqn:?. + destruct (List.existsb pred (ATree.elements e)) eqn:?. (* yes *) rewrite List.existsb_exists in Heqb. destruct Heqb as [[id [b sz]] [A B]]. @@ -403,14 +403,14 @@ Proof. destruct (eq_block sp sp'); try discriminate. destruct (andb_prop _ _ B). left. apply is_reachable_intro with id b sz delta. - apply PTree.elements_complete; auto. + apply ATree.elements_complete; auto. congruence. split; eapply proj_sumbool_true; eauto. (* no *) right; red; intro NE; inv NE. - assert (existsb pred (PTree.elements e) = true). + assert (existsb pred (ATree.elements e) = true). rewrite List.existsb_exists. exists (id, (b, sz)); split. - apply PTree.elements_correct; auto. + apply ATree.elements_correct; auto. simpl. rewrite H0. rewrite dec_eq_true. unfold proj_sumbool. destruct H1. rewrite zle_true; auto. rewrite zlt_true; auto. congruence. @@ -548,7 +548,7 @@ Lemma match_callstack_set_temp: forall f cenv e le te sp lo hi cs bound tbound m tm tf id v tv, Val.inject f v tv -> match_callstack f m tm (Frame cenv tf e le te sp lo hi :: cs) bound tbound -> - match_callstack f m tm (Frame cenv tf e (PTree.set id v le) (PTree.set id tv te) sp lo hi :: cs) bound tbound. + match_callstack f m tm (Frame cenv tf e (ATree.set id v le) (ATree.set id tv te) sp lo hi :: cs) bound tbound. Proof. intros. inv H0. constructor; auto. eapply match_temps_assign; eauto. @@ -560,22 +560,22 @@ Qed. Lemma in_blocks_of_env: forall e id b sz, - e!id = Some(b, sz) -> In (b, 0, sz) (blocks_of_env e). + e$id = Some(b, sz) -> In (b, 0, sz) (blocks_of_env e). Proof. unfold blocks_of_env; intros. change (b, 0, sz) with (block_of_binding (id, (b, sz))). - apply List.in_map. apply PTree.elements_correct. auto. + apply List.in_map. apply ATree.elements_correct. auto. Qed. Lemma in_blocks_of_env_inv: forall b lo hi e, In (b, lo, hi) (blocks_of_env e) -> - exists id, e!id = Some(b, hi) /\ lo = 0. + exists id, e$id = Some(b, hi) /\ lo = 0. Proof. unfold blocks_of_env; intros. exploit list_in_map_inv; eauto. intros [[id [b' sz]] [A B]]. unfold block_of_binding in A. inv A. - exists id; intuition. apply PTree.elements_complete. auto. + exists id; intuition. apply ATree.elements_complete. auto. Qed. Lemma match_callstack_freelist: @@ -674,7 +674,7 @@ Lemma match_callstack_alloc_right: Mem.alloc tm 0 tf.(fn_stackspace) = (tm', sp) -> Mem.inject f m tm -> match_temps f le te -> - (forall id, cenv!id = None) -> + (forall id, cenv$id = None) -> match_callstack f m tm' (Frame cenv tf empty_env le te sp (Mem.nextblock m) (Mem.nextblock m) :: cs) (Mem.nextblock m) (Mem.nextblock tm'). @@ -687,12 +687,12 @@ Proof. unfold block in *; xomega. auto. constructor; intros. - rewrite H3. rewrite PTree.gempty. constructor. + rewrite H3. rewrite ATree.gempty. constructor. xomega. - rewrite PTree.gempty in H4; discriminate. + rewrite ATree.gempty in H4; discriminate. eelim Mem.fresh_block_alloc; eauto. eapply Mem.valid_block_inject_2; eauto. rewrite RES. change (Mem.valid_block tm tb). eapply Mem.valid_block_inject_2; eauto. - red; intros. rewrite PTree.gempty in H4. discriminate. + red; intros. rewrite ATree.gempty in H4. discriminate. red; intros. left. eapply Mem.perm_alloc_2; eauto. eapply match_callstack_invariant with (tm1 := tm); eauto. rewrite RES; auto. @@ -702,16 +702,16 @@ Qed. Lemma match_callstack_alloc_left: forall f1 m1 tm id cenv tf e le te sp lo cs sz m2 b f2 ofs, match_callstack f1 m1 tm - (Frame (PTree.remove id cenv) tf e le te sp lo (Mem.nextblock m1) :: cs) + (Frame (ATree.remove id cenv) tf e le te sp lo (Mem.nextblock m1) :: cs) (Mem.nextblock m1) (Mem.nextblock tm) -> Mem.alloc m1 0 sz = (m2, b) -> - cenv!id = Some ofs -> + cenv$id = Some ofs -> inject_incr f1 f2 -> f2 b = Some(sp, ofs) -> (forall b', b' <> b -> f2 b' = f1 b') -> - e!id = None -> + e$id = None -> match_callstack f2 m2 tm - (Frame cenv tf (PTree.set id (b, sz) e) le te sp lo (Mem.nextblock m2) :: cs) + (Frame cenv tf (ATree.set id (b, sz) e) le te sp lo (Mem.nextblock m2) :: cs) (Mem.nextblock m2) (Mem.nextblock tm). Proof. intros. inv H. @@ -723,13 +723,13 @@ Proof. auto. eapply match_temps_invariant; eauto. eapply match_env_alloc; eauto. - red; intros. rewrite PTree.gsspec in H. destruct (peq id0 id). + red; intros. rewrite ATree.gsspec in H. destruct (ATree.elt_eq id0 id). inversion H. subst b0 sz0 id0. eapply Mem.perm_alloc_3; eauto. eapply BOUND0; eauto. eapply Mem.perm_alloc_4; eauto. exploit me_bounded; eauto. unfold block in *; xomega. red; intros. exploit PERM; eauto. intros [A|A]. auto. right. inv A. apply is_reachable_intro with id0 b0 sz0 delta; auto. - rewrite PTree.gso. auto. congruence. + rewrite ATree.gso. auto. congruence. eapply match_callstack_invariant with (m1 := m1); eauto. intros. eapply Mem.perm_alloc_4; eauto. unfold block in *; xomega. @@ -745,26 +745,26 @@ Qed. local variables as sub-blocks of the Cminor stack data. This is the most difficult part of the proof. *) Definition cenv_remove (cenv: compilenv) (vars: list (ident * Z)) : compilenv := - fold_right (fun id_lv ce => PTree.remove (fst id_lv) ce) cenv vars. + fold_right (fun id_lv ce => ATree.remove (fst id_lv) ce) cenv vars. Remark cenv_remove_gso: forall id vars cenv, ~In id (map fst vars) -> - PTree.get id (cenv_remove cenv vars) = PTree.get id cenv. + ATree.get id (cenv_remove cenv vars) = ATree.get id cenv. Proof. induction vars; simpl; intros. auto. - rewrite PTree.gro. apply IHvars. intuition. intuition. + rewrite ATree.gro. apply IHvars. intuition. intuition. Qed. Remark cenv_remove_gss: forall id vars cenv, In id (map fst vars) -> - PTree.get id (cenv_remove cenv vars) = None. + ATree.get id (cenv_remove cenv vars) = None. Proof. induction vars; simpl; intros. contradiction. - rewrite PTree.grspec. destruct (PTree.elt_eq id (fst a)). auto. + rewrite ATree.grspec. destruct (ATree.elt_eq id (fst a)). auto. destruct H. intuition. eauto. Qed. @@ -772,7 +772,7 @@ Definition cenv_compat (cenv: compilenv) (vars: list (ident * Z)) (tsz: Z) : Pro forall id sz, In (id, sz) vars -> exists ofs, - PTree.get id cenv = Some ofs + ATree.get id cenv = Some ofs /\ Mem.inj_offset_aligned ofs sz /\ 0 <= ofs /\ ofs + Z.max 0 sz <= tsz. @@ -780,13 +780,13 @@ Definition cenv_compat (cenv: compilenv) (vars: list (ident * Z)) (tsz: Z) : Pro Definition cenv_separated (cenv: compilenv) (vars: list (ident * Z)) : Prop := forall id1 sz1 ofs1 id2 sz2 ofs2, In (id1, sz1) vars -> In (id2, sz2) vars -> - PTree.get id1 cenv = Some ofs1 -> PTree.get id2 cenv = Some ofs2 -> + ATree.get id1 cenv = Some ofs1 -> ATree.get id2 cenv = Some ofs2 -> id1 <> id2 -> ofs1 + sz1 <= ofs2 \/ ofs2 + sz2 <= ofs1. Definition cenv_mem_separated (cenv: compilenv) (vars: list (ident * Z)) (f: meminj) (sp: block) (m: mem) : Prop := forall id sz ofs b delta ofs' k p, - In (id, sz) vars -> PTree.get id cenv = Some ofs -> + In (id, sz) vars -> ATree.get id cenv = Some ofs -> f b = Some (sp, delta) -> Mem.perm m b ofs' k p -> ofs <= ofs' + delta < sz + ofs -> False. @@ -804,7 +804,7 @@ Lemma match_callstack_alloc_variables_rec: cenv_compat cenv vars (fn_stackspace tf) -> cenv_separated cenv vars -> cenv_mem_separated cenv vars f1 sp m1 -> - (forall id sz, In (id, sz) vars -> e1!id = None) -> + (forall id sz, In (id, sz) vars -> e1$id = None) -> match_callstack f1 m1 tm (Frame (cenv_remove cenv vars) tf e1 le te sp lo (Mem.nextblock m1) :: cs) (Mem.nextblock m1) (Mem.nextblock tm) -> @@ -844,7 +844,7 @@ Proof. omega. eapply SEP2. apply in_cons; eauto. eauto. rewrite D in H5; eauto. eauto. auto. - intros. rewrite PTree.gso. eapply UNBOUND; eauto with coqlib. + intros. rewrite ATree.gso. eapply UNBOUND; eauto with coqlib. red; intros; subst id0. elim H3. change id with (fst (id, sz0)). apply in_map; auto. eapply match_callstack_alloc_left; eauto. rewrite cenv_remove_gso; auto. @@ -859,7 +859,7 @@ Lemma match_callstack_alloc_variables: list_norepet (map fst vars) -> cenv_compat cenv vars (fn_stackspace fn) -> cenv_separated cenv vars -> - (forall id ofs, cenv!id = Some ofs -> In id (map fst vars)) -> + (forall id ofs, cenv$id = Some ofs -> In id (map fst vars)) -> Mem.inject f1 m1 tm1 -> match_callstack f1 m1 tm1 cs (Mem.nextblock m1) (Mem.nextblock tm1) -> match_temps f1 le te -> @@ -875,12 +875,12 @@ Proof. intros. apply Mem.perm_implies with Freeable; auto with mem. eapply Mem.perm_alloc_2; eauto. instantiate (1 := f1). red; intros. eelim Mem.fresh_block_alloc; eauto. eapply Mem.valid_block_inject_2; eauto. - intros. apply PTree.gempty. + intros. apply ATree.gempty. eapply match_callstack_alloc_right; eauto. - intros. destruct (In_dec peq id (map fst vars)). + intros. destruct (In_dec ident_eq id (map fst vars)). apply cenv_remove_gss; auto. rewrite cenv_remove_gso; auto. - destruct (cenv!id) as [ofs|] eqn:?; auto. elim n; eauto. + destruct (cenv$id) as [ofs|] eqn:?; auto. elim n; eauto. eapply Mem.alloc_right_inject; eauto. Qed. @@ -975,21 +975,21 @@ Proof. apply EITHER in H. destruct H as [[P Q] | P]. exploit COMPAT; eauto. intros [ofs [A [B [C D]]]]. exists ofs. - split. rewrite PTree.gso; auto. + split. rewrite ATree.gso; auto. split. auto. split. auto. zify; omega. inv P. exists (align sz1 (block_alignment sz)). - split. apply PTree.gss. + split. apply ATree.gss. split. apply inj_offset_aligned_block. split. omega. omega. apply EITHER in H; apply EITHER in H0. destruct H as [[P Q] | P]; destruct H0 as [[R S] | R]. - rewrite PTree.gso in *; auto. eapply SEP; eauto. - inv R. rewrite PTree.gso in H1; auto. rewrite PTree.gss in H2; inv H2. + rewrite ATree.gso in *; auto. eapply SEP; eauto. + inv R. rewrite ATree.gso in H1; auto. rewrite ATree.gss in H2; inv H2. exploit COMPAT; eauto. intros [ofs [A [B [C D]]]]. assert (ofs = ofs1) by congruence. subst ofs. left. zify; omega. - inv P. rewrite PTree.gso in H2; auto. rewrite PTree.gss in H1; inv H1. + inv P. rewrite ATree.gso in H2; auto. rewrite ATree.gss in H1; inv H1. exploit COMPAT; eauto. intros [ofs [A [B [C D]]]]. assert (ofs = ofs2) by congruence. subst ofs. right. zify; omega. @@ -1064,28 +1064,28 @@ Qed. Lemma assign_variables_domain: forall id vars cesz, - (fst (assign_variables cesz vars))!id <> None -> - (fst cesz)!id <> None \/ In id (map fst vars). + (fst (assign_variables cesz vars))$id <> None -> + (fst cesz)$id <> None \/ In id (map fst vars). Proof. induction vars; simpl; intros. auto. exploit IHvars; eauto. unfold assign_variable. destruct a as [id1 sz1]. destruct cesz as [cenv stksz]. simpl. - rewrite PTree.gsspec. destruct (peq id id1). auto. tauto. + rewrite ATree.gsspec. destruct (ATree.elt_eq id id1). auto. tauto. Qed. Lemma build_compilenv_domain: forall f cenv sz id ofs, build_compilenv f = (cenv, sz) -> - cenv!id = Some ofs -> In id (map fst (Csharpminor.fn_vars f)). + cenv$id = Some ofs -> In id (map fst (Csharpminor.fn_vars f)). Proof. unfold build_compilenv; intros. set (vars1 := Csharpminor.fn_vars f) in *. generalize (VarSort.Permuted_sort vars1). intros P. set (vars2 := VarSort.sort vars1) in *. - generalize (assign_variables_domain id vars2 (PTree.empty Z, 0)). + generalize (assign_variables_domain id vars2 (ATree.empty Z, 0)). rewrite H. simpl. intros. destruct H1. congruence. - rewrite PTree.gempty in H1. congruence. + rewrite ATree.gempty in H1. congruence. apply Permutation_in with (map fst vars2); auto. apply Permutation_map. apply Permutation_sym; auto. Qed. @@ -1093,19 +1093,19 @@ Qed. (** Initialization of C#minor temporaries and Cminor local variables. *) Lemma create_undef_temps_val: - forall id v temps, (create_undef_temps temps)!id = Some v -> In id temps /\ v = Vundef. + forall id v temps, (create_undef_temps temps)$id = Some v -> In id temps /\ v = Vundef. Proof. induction temps; simpl; intros. - rewrite PTree.gempty in H. congruence. - rewrite PTree.gsspec in H. destruct (peq id a). + rewrite ATree.gempty in H. congruence. + rewrite ATree.gsspec in H. destruct (ATree.elt_eq id a). split. auto. congruence. exploit IHtemps; eauto. tauto. Qed. Fixpoint set_params' (vl: list val) (il: list ident) (te: Cminor.env) : Cminor.env := match il, vl with - | i1 :: is, v1 :: vs => set_params' vs is (PTree.set i1 v1 te) - | i1 :: is, nil => set_params' nil is (PTree.set i1 Vundef te) + | i1 :: is, v1 :: vs => set_params' vs is (ATree.set i1 v1 te) + | i1 :: is, nil => set_params' nil is (ATree.set i1 Vundef te) | _, _ => te end. @@ -1116,88 +1116,88 @@ Lemma bind_parameters_agree_rec: match_temps f le1 te -> match_temps f le2 (set_params' tvals vars te). Proof. -Opaque PTree.set. +Opaque ATree.set. induction vars; simpl; intros. destruct vals; try discriminate. inv H. auto. destruct vals; try discriminate. inv H0. simpl. eapply IHvars; eauto. - red; intros. rewrite PTree.gsspec in *. destruct (peq id a). + red; intros. rewrite ATree.gsspec in *. destruct (ATree.elt_eq id a). inv H0. exists v'; auto. apply H1; auto. Qed. Lemma set_params'_outside: - forall id il vl te, ~In id il -> (set_params' vl il te)!id = te!id. + forall id il vl te, ~In id il -> (set_params' vl il te)$id = te$id. Proof. induction il; simpl; intros. auto. destruct vl; rewrite IHil. - apply PTree.gso. intuition. intuition. - apply PTree.gso. intuition. intuition. + apply ATree.gso. intuition. intuition. + apply ATree.gso. intuition. intuition. Qed. Lemma set_params'_inside: forall id il vl te1 te2, In id il -> - (set_params' vl il te1)!id = (set_params' vl il te2)!id. + (set_params' vl il te1)$id = (set_params' vl il te2)$id. Proof. induction il; simpl; intros. contradiction. - destruct vl; destruct (List.in_dec peq id il); auto; + destruct vl; destruct (List.in_dec ident_eq id il); auto; repeat rewrite set_params'_outside; auto; - assert (a = id) by intuition; subst a; repeat rewrite PTree.gss; auto. + assert (a = id) by intuition; subst a; repeat rewrite ATree.gss; auto. Qed. Lemma set_params_set_params': forall il vl id, list_norepet il -> - (set_params vl il)!id = (set_params' vl il (PTree.empty val))!id. + (set_params vl il)$id = (set_params' vl il (ATree.empty val))$id. Proof. induction il; simpl; intros. auto. inv H. destruct vl. - rewrite PTree.gsspec. destruct (peq id a). - subst a. rewrite set_params'_outside; auto. rewrite PTree.gss; auto. + rewrite ATree.gsspec. destruct (ATree.elt_eq id a). + subst a. rewrite set_params'_outside; auto. rewrite ATree.gss; auto. rewrite IHil; auto. - destruct (List.in_dec peq id il). apply set_params'_inside; auto. - repeat rewrite set_params'_outside; auto. rewrite PTree.gso; auto. - rewrite PTree.gsspec. destruct (peq id a). - subst a. rewrite set_params'_outside; auto. rewrite PTree.gss; auto. + destruct (List.in_dec ident_eq id il). apply set_params'_inside; auto. + repeat rewrite set_params'_outside; auto. rewrite ATree.gso; auto. + rewrite ATree.gsspec. destruct (ATree.elt_eq id a). + subst a. rewrite set_params'_outside; auto. rewrite ATree.gss; auto. rewrite IHil; auto. - destruct (List.in_dec peq id il). apply set_params'_inside; auto. - repeat rewrite set_params'_outside; auto. rewrite PTree.gso; auto. + destruct (List.in_dec ident_eq id il). apply set_params'_inside; auto. + repeat rewrite set_params'_outside; auto. rewrite ATree.gso; auto. Qed. Lemma set_locals_outside: forall e id il, - ~In id il -> (set_locals il e)!id = e!id. + ~In id il -> (set_locals il e)$id = e$id. Proof. induction il; simpl; intros. auto. - rewrite PTree.gso. apply IHil. tauto. intuition. + rewrite ATree.gso. apply IHil. tauto. intuition. Qed. Lemma set_locals_inside: forall e id il, - In id il -> (set_locals il e)!id = Some Vundef. + In id il -> (set_locals il e)$id = Some Vundef. Proof. induction il; simpl; intros. contradiction. - destruct H. subst a. apply PTree.gss. - rewrite PTree.gsspec. destruct (peq id a). auto. auto. + destruct H. subst a. apply ATree.gss. + rewrite ATree.gsspec. destruct (ATree.elt_eq id a). auto. auto. Qed. Lemma set_locals_set_params': forall vars vals params id, list_norepet params -> list_disjoint params vars -> - (set_locals vars (set_params vals params)) ! id = - (set_params' vals params (set_locals vars (PTree.empty val))) ! id. + (set_locals vars (set_params vals params)) $ id = + (set_params' vals params (set_locals vars (ATree.empty val))) $ id. Proof. - intros. destruct (in_dec peq id vars). + intros. destruct (in_dec ident_eq id vars). assert (~In id params). apply list_disjoint_notin with vars; auto. apply list_disjoint_sym; auto. rewrite set_locals_inside; auto. rewrite set_params'_outside; auto. rewrite set_locals_inside; auto. rewrite set_locals_outside; auto. rewrite set_params_set_params'; auto. - destruct (in_dec peq id params). + destruct (in_dec ident_eq id params). apply set_params'_inside; auto. repeat rewrite set_params'_outside; auto. rewrite set_locals_outside; auto. @@ -1213,7 +1213,7 @@ Lemma bind_parameters_agree: Proof. intros; red; intros. exploit bind_parameters_agree_rec; eauto. - instantiate (1 := set_locals temps (PTree.empty val)). + instantiate (1 := set_locals temps (ATree.empty val)). red; intros. exploit create_undef_temps_val; eauto. intros [A B]. subst v0. exists Vundef; split. apply set_locals_inside; auto. auto. intros [v' [A B]]. exists v'; split; auto. @@ -1434,7 +1434,7 @@ Lemma var_addr_correct: /\ Val.inject f (Vptr b Ptrofs.zero) tv. Proof. unfold var_addr; intros. - assert (match_var f sp e!id cenv!id). + assert (match_var f sp e$id cenv$id). inv H. inv MENV. auto. inv H1; inv H0; try congruence. (* local *) @@ -2043,7 +2043,7 @@ Proof. eapply external_call_nextblock; eauto. eapply external_call_nextblock; eauto. econstructor; eauto. -Opaque PTree.set. +Opaque ATree.set. unfold set_optvar. destruct optid; simpl. eapply match_callstack_set_temp; eauto. auto. @@ -2232,7 +2232,7 @@ Proof. eapply match_program_main; eauto. eexact FIND. rewrite <- H2. apply sig_preserved; auto. - eapply match_callstate with (f := Mem.flat_inj (Mem.nextblock m0)) (cs := @nil frame) (cenv := PTree.empty Z). + eapply match_callstate with (f := Mem.flat_inj (Mem.nextblock m0)) (cs := @nil frame) (cenv := ATree.empty Z). auto. eapply Genv.initmem_inject; eauto. apply mcs_nil with (Mem.nextblock m0). apply match_globalenvs_init; auto. xomega. xomega. diff --git a/cfrontend/Csem.v b/cfrontend/Csem.v index 0c3e00de73..66606e953a 100644 --- a/cfrontend/Csem.v +++ b/cfrontend/Csem.v @@ -46,9 +46,9 @@ Definition globalenv (p: program) := The current value of the variable is stored in the associated memory block. *) -Definition env := PTree.t (block * type). (* map variable -> location & type *) +Definition env := ATree.t (block * type). (* map variable -> location & type *) -Definition empty_env: env := (PTree.empty (block * type)). +Definition empty_env: env := (ATree.empty (block * type)). Section SEMANTICS. @@ -125,7 +125,7 @@ Inductive alloc_variables: env -> mem -> | alloc_variables_cons: forall e m id ty vars m1 b1 m2 e2, Mem.alloc m 0 (sizeof ge ty) = (m1, b1) -> - alloc_variables (PTree.set id (b1, ty) e) m1 vars e2 m2 -> + alloc_variables (ATree.set id (b1, ty) e) m1 vars e2 m2 -> alloc_variables e m ((id, ty) :: vars) e2 m2. (** Initialization of local variables that are parameters to a function. @@ -141,7 +141,7 @@ Inductive bind_parameters (e: env): bind_parameters e m nil nil m | bind_parameters_cons: forall m id ty params v1 vl b m1 m2, - PTree.get id e = Some(b, ty) -> + ATree.get id e = Some(b, ty) -> assign_loc ty m b Ptrofs.zero v1 E0 m1 -> bind_parameters e m1 params vl m2 -> bind_parameters e m ((id, ty) :: params) (v1 :: vl) m2. @@ -152,7 +152,7 @@ Definition block_of_binding (id_b_ty: ident * (block * type)) := match id_b_ty with (id, (b, ty)) => (b, 0, sizeof ge ty) end. Definition blocks_of_env (e: env) : list (block * Z * Z) := - List.map block_of_binding (PTree.elements e). + List.map block_of_binding (ATree.elements e). (** Selection of the appropriate case of a [switch], given the value [n] of the selector expression. *) @@ -209,11 +209,11 @@ Variable e: env. Inductive lred: expr -> mem -> expr -> mem -> Prop := | red_var_local: forall x ty m b, - e!x = Some(b, ty) -> + e$x = Some(b, ty) -> lred (Evar x ty) m (Eloc b Ptrofs.zero ty) m | red_var_global: forall x ty m b, - e!x = None -> + e$x = None -> Genv.find_symbol ge x = Some b -> lred (Evar x ty) m (Eloc b Ptrofs.zero ty) m @@ -221,12 +221,12 @@ Inductive lred: expr -> mem -> expr -> mem -> Prop := lred (Ederef (Eval (Vptr b ofs) ty1) ty) m (Eloc b ofs ty) m | red_field_struct: forall b ofs id co a f ty m delta, - ge.(genv_cenv)!id = Some co -> + ge.(genv_cenv)$id = Some co -> field_offset ge f (co_members co) = OK delta -> lred (Efield (Eval (Vptr b ofs) (Tstruct id a)) f ty) m (Eloc b (Ptrofs.add ofs (Ptrofs.repr delta)) ty) m | red_field_union: forall b ofs id co a f ty m, - ge.(genv_cenv)!id = Some co -> + ge.(genv_cenv)$id = Some co -> lred (Efield (Eval (Vptr b ofs) (Tunion id a)) f ty) m (Eloc b ofs ty) m. diff --git a/cfrontend/Csharpminor.v b/cfrontend/Csharpminor.v index 4393640ced..1c924820b2 100644 --- a/cfrontend/Csharpminor.v +++ b/cfrontend/Csharpminor.v @@ -111,18 +111,18 @@ Definition funsig (fd: fundef) := *) Definition genv := Genv.t fundef unit. -Definition env := PTree.t (block * Z). -Definition temp_env := PTree.t val. +Definition env := ATree.t (block * Z). +Definition temp_env := ATree.t val. -Definition empty_env : env := PTree.empty (block * Z). -Definition empty_temp_env : temp_env := PTree.empty val. +Definition empty_env : env := ATree.empty (block * Z). +Definition empty_temp_env : temp_env := ATree.empty val. (** Initialization of temporary variables *) Fixpoint create_undef_temps (temps: list ident) : temp_env := match temps with - | nil => PTree.empty val - | id :: temps' => PTree.set id Vundef (create_undef_temps temps') + | nil => ATree.empty val + | id :: temps' => ATree.set id Vundef (create_undef_temps temps') end. (** Initialization of temporaries that are parameters. *) @@ -131,7 +131,7 @@ Fixpoint bind_parameters (formals: list ident) (args: list val) (le: temp_env) : option temp_env := match formals, args with | nil, nil => Some le - | id :: xl, v :: vl => bind_parameters xl vl (PTree.set id v le) + | id :: xl, v :: vl => bind_parameters xl vl (ATree.set id v le) | _, _ => None end. @@ -275,7 +275,7 @@ Inductive alloc_variables: env -> mem -> | alloc_variables_cons: forall e m id sz vars m1 b1 m2 e2, Mem.alloc m 0 sz = (m1, b1) -> - alloc_variables (PTree.set id (b1, sz) e) m1 vars e2 m2 -> + alloc_variables (ATree.set id (b1, sz) e) m1 vars e2 m2 -> alloc_variables e m ((id, sz) :: vars) e2 m2. (** List of blocks mentioned in an environment, with low and high bounds *) @@ -284,7 +284,7 @@ Definition block_of_binding (id_b_sz: ident * (block * Z)) := match id_b_sz with (id, (b, sz)) => (b, 0, sz) end. Definition blocks_of_env (e: env) : list (block * Z * Z) := - List.map block_of_binding (PTree.elements e). + List.map block_of_binding (ATree.elements e). Section RELSEM. @@ -297,11 +297,11 @@ Variable ge: genv. Inductive eval_var_addr: env -> ident -> block -> Prop := | eval_var_addr_local: forall e id b sz, - PTree.get id e = Some (b, sz) -> + ATree.get id e = Some (b, sz) -> eval_var_addr e id b | eval_var_addr_global: forall e id b, - PTree.get id e = None -> + ATree.get id e = None -> Genv.find_symbol ge id = Some b -> eval_var_addr e id b. @@ -317,7 +317,7 @@ Variable m: mem. Inductive eval_expr: expr -> val -> Prop := | eval_Evar: forall id v, - le!id = Some v -> + le$id = Some v -> eval_expr (Evar id) v | eval_Eaddrof: forall id b, eval_var_addr e id b -> @@ -373,7 +373,7 @@ Inductive step: state -> trace -> state -> Prop := | step_set: forall f id a k e le m v, eval_expr e le m a v -> step (State f (Sset id a) k e le m) - E0 (State f Sskip k e (PTree.set id v le) m) + E0 (State f Sskip k e (ATree.set id v le) m) | step_store: forall f chunk addr a k e le m vaddr v m', eval_expr e le m addr vaddr -> diff --git a/cfrontend/Cshmgen.v b/cfrontend/Cshmgen.v index 792a73f98d..6c46a6c823 100644 --- a/cfrontend/Cshmgen.v +++ b/cfrontend/Cshmgen.v @@ -443,7 +443,7 @@ Definition transl_binop (ce: composite_env) Definition make_field_access (ce: composite_env) (ty: type) (f: ident) (a: expr) : res expr := match ty with | Tstruct id _ => - match ce!id with + match ce$id with | None => Error (MSG "Undefined struct " :: CTX id :: nil) | Some co => diff --git a/cfrontend/Cshmgenproof.v b/cfrontend/Cshmgenproof.v index 09e31cb2b5..c1a3fdceac 100644 --- a/cfrontend/Cshmgenproof.v +++ b/cfrontend/Cshmgenproof.v @@ -112,14 +112,14 @@ Proof. split. - symmetry. apply Ctypes.sizeof_stable; auto. - revert C. induction t; simpl; auto; - destruct (prog_comp_env cunit)!i as [co|] eqn:X; try discriminate; erewrite H1 by eauto; auto. + destruct (prog_comp_env cunit)$i as [co|] eqn:X; try discriminate; erewrite H1 by eauto; auto. Qed. Lemma field_offset_stable: forall (cunit prog: Clight.program) id co f, linkorder cunit prog -> - cunit.(prog_comp_env)!id = Some co -> - prog.(prog_comp_env)!id = Some co /\ + cunit.(prog_comp_env)$id = Some co -> + prog.(prog_comp_env)$id = Some co /\ field_offset prog.(prog_comp_env) f (co_members co) = field_offset cunit.(prog_comp_env) f (co_members co). Proof. intros. @@ -1038,19 +1038,19 @@ Record match_env (e: Clight.env) (te: Csharpminor.env) : Prop := mk_match_env { me_local: forall id b ty, - e!id = Some (b, ty) -> te!id = Some(b, Ctypes.sizeof ge ty); + e$id = Some (b, ty) -> te$id = Some(b, Ctypes.sizeof ge ty); me_local_inv: forall id b sz, - te!id = Some (b, sz) -> exists ty, e!id = Some(b, ty) + te$id = Some (b, sz) -> exists ty, e$id = Some(b, ty) }. Lemma match_env_globals: forall e te id, match_env e te -> - e!id = None -> - te!id = None. + e$id = None -> + te$id = None. Proof. - intros. destruct (te!id) as [[b sz] | ] eqn:?; auto. + intros. destruct (te$id) as [[b sz] | ] eqn:?; auto. exploit me_local_inv; eauto. intros [ty EQ]. congruence. Qed. @@ -1066,8 +1066,8 @@ Proof. end). assert (list_forall2 (fun i_x i_y => fst i_x = fst i_y /\ R (snd i_x) (snd i_y)) - (PTree.elements e) (PTree.elements te)). - apply PTree.elements_canonical_order. + (ATree.elements e) (ATree.elements te)). + apply ATree.elements_canonical_order. intros id [b ty] GET. exists (b, Ctypes.sizeof ge ty); split. eapply me_local; eauto. red; auto. intros id [b sz] GET. exploit me_local_inv; eauto. intros [ty EQ]. exploit me_local; eauto. intros EQ1. @@ -1095,8 +1095,8 @@ Lemma match_env_empty: Proof. unfold Clight.empty_env, Csharpminor.empty_env. constructor. - intros until ty. repeat rewrite PTree.gempty. congruence. - intros until sz. rewrite PTree.gempty. congruence. + intros until ty. repeat rewrite ATree.gempty. congruence. + intros until sz. rewrite ATree.gempty. congruence. Qed. (** The following lemmas establish the [match_env] invariant at @@ -1117,15 +1117,15 @@ Proof. - inv H0. exists te1; split. constructor. auto. - monadInv H2. monadInv EQ. simpl in *. exploit transl_sizeof. eexact H. eauto. intros SZ; rewrite SZ. - exploit (IHalloc_variables x0 (PTree.set id (b1, Ctypes.sizeof ge ty) te1)). + exploit (IHalloc_variables x0 (ATree.set id (b1, Ctypes.sizeof ge ty) te1)). auto. constructor. (* me_local *) - intros until ty0. repeat rewrite PTree.gsspec. - destruct (peq id0 id); intros. congruence. eapply me_local; eauto. + intros until ty0. repeat rewrite ATree.gsspec. + destruct (ATree.elt_eq id0 id); intros. congruence. eapply me_local; eauto. (* me_local_inv *) - intros until sz. repeat rewrite PTree.gsspec. - destruct (peq id0 id); intros. exists ty; congruence. eapply me_local_inv; eauto. + intros until sz. repeat rewrite ATree.gsspec. + destruct (ATree.elt_eq id0 id); intros. exists ty; congruence. eapply me_local_inv; eauto. intros [te2 [ALLOC MENV]]. exists te2; split. econstructor; eauto. auto. Qed. @@ -1239,7 +1239,7 @@ Proof. simpl in TR. eauto. - (* field struct *) unfold make_field_access in EQ0. rewrite H1 in EQ0. - destruct (prog_comp_env cunit)!id as [co'|] eqn:CO; monadInv EQ0. + destruct (prog_comp_env cunit)$id as [co'|] eqn:CO; monadInv EQ0. exploit field_offset_stable. eexact LINK. eauto. instantiate (1 := i). intros [A B]. rewrite <- B in EQ1. assert (x0 = delta) by (unfold ge in *; simpl in *; congruence). diff --git a/cfrontend/Cstrategy.v b/cfrontend/Cstrategy.v index 2f73106833..ecb88bf905 100644 --- a/cfrontend/Cstrategy.v +++ b/cfrontend/Cstrategy.v @@ -88,10 +88,10 @@ Inductive eval_simple_lvalue: expr -> block -> ptrofs -> Prop := | esl_loc: forall b ofs ty, eval_simple_lvalue (Eloc b ofs ty) b ofs | esl_var_local: forall x ty b, - e!x = Some(b, ty) -> + e$x = Some(b, ty) -> eval_simple_lvalue (Evar x ty) b Ptrofs.zero | esl_var_global: forall x ty b, - e!x = None -> + e$x = None -> Genv.find_symbol ge x = Some b -> eval_simple_lvalue (Evar x ty) b Ptrofs.zero | esl_deref: forall r ty b ofs, @@ -100,13 +100,13 @@ Inductive eval_simple_lvalue: expr -> block -> ptrofs -> Prop := | esl_field_struct: forall r f ty b ofs id co a delta, eval_simple_rvalue r (Vptr b ofs) -> typeof r = Tstruct id a -> - ge.(genv_cenv)!id = Some co -> + ge.(genv_cenv)$id = Some co -> field_offset ge f (co_members co) = OK delta -> eval_simple_lvalue (Efield r f ty) b (Ptrofs.add ofs (Ptrofs.repr delta)) | esl_field_union: forall r f ty b ofs id co a, eval_simple_rvalue r (Vptr b ofs) -> typeof r = Tunion id a -> - ge.(genv_cenv)!id = Some co -> + ge.(genv_cenv)$id = Some co -> eval_simple_lvalue (Efield r f ty) b ofs with eval_simple_rvalue: expr -> val -> Prop := @@ -521,15 +521,15 @@ Definition invert_expr_prop (a: expr) (m: mem) : Prop := | Eloc b ofs ty => False | Evar x ty => exists b, - e!x = Some(b, ty) - \/ (e!x = None /\ Genv.find_symbol ge x = Some b) + e$x = Some(b, ty) + \/ (e$x = None /\ Genv.find_symbol ge x = Some b) | Ederef (Eval v ty1) ty => exists b, exists ofs, v = Vptr b ofs | Efield (Eval v ty1) f ty => exists b, exists ofs, v = Vptr b ofs /\ match ty1 with - | Tstruct id _ => exists co delta, ge.(genv_cenv)!id = Some co /\ field_offset ge f (co_members co) = Errors.OK delta - | Tunion id _ => exists co, ge.(genv_cenv)!id = Some co + | Tstruct id _ => exists co delta, ge.(genv_cenv)$id = Some co /\ field_offset ge f (co_members co) = Errors.OK delta + | Tunion id _ => exists co, ge.(genv_cenv)$id = Some co | _ => False end | Eval v ty => False diff --git a/cfrontend/Ctypes.v b/cfrontend/Ctypes.v index 036b768bca..8805955c63 100644 --- a/cfrontend/Ctypes.v +++ b/cfrontend/Ctypes.v @@ -190,7 +190,7 @@ Record composite : Type := { co_sizeof_alignof: (co_alignof | co_sizeof) }. -Definition composite_env : Type := PTree.t composite. +Definition composite_env : Type := ATree.t composite. (** * Operations over types *) @@ -240,7 +240,7 @@ Fixpoint complete_type (env: composite_env) (t: type) : bool := | Tarray t' _ _ => complete_type env t' | Tfunction _ _ _ => false | Tstruct id _ | Tunion id _ => - match env!id with Some co => true | None => false end + match env$id with Some co => true | None => false end end. Definition complete_or_function_type (env: composite_env) (t: type) : bool := @@ -280,7 +280,7 @@ Fixpoint alignof (env: composite_env) (t: type) : Z := | Tarray t' _ _ => alignof env t' | Tfunction _ _ _ => 1 | Tstruct id _ | Tunion id _ => - match env!id with Some co => co_alignof co | None => 1 end + match env$id with Some co => co_alignof co | None => 1 end end). Remark align_attr_two_p: @@ -310,8 +310,8 @@ Proof. exists (if Archi.ptr64 then 3%nat else 2%nat); destruct Archi.ptr64; auto. apply IHt. exists 0%nat; auto. - destruct (env!i). apply co_alignof_two_p. exists 0%nat; auto. - destruct (env!i). apply co_alignof_two_p. exists 0%nat; auto. + destruct (env$i). apply co_alignof_two_p. exists 0%nat; auto. + destruct (env$i). apply co_alignof_two_p. exists 0%nat; auto. Qed. Lemma alignof_pos: @@ -343,7 +343,7 @@ Fixpoint sizeof (env: composite_env) (t: type) : Z := | Tarray t' n _ => sizeof env t' * Z.max 0 n | Tfunction _ _ _ => 1 | Tstruct id _ | Tunion id _ => - match env!id with Some co => co_sizeof co | None => 0 end + match env$id with Some co => co_sizeof co | None => 0 end end. Lemma sizeof_pos: @@ -354,8 +354,8 @@ Proof. destruct f; omega. destruct Archi.ptr64; omega. change 0 with (0 * Z.max 0 z) at 2. apply Zmult_ge_compat_r. auto. xomega. - destruct (env!i). apply co_sizeof_pos. omega. - destruct (env!i). apply co_sizeof_pos. omega. + destruct (env$i). apply co_sizeof_pos. omega. + destruct (env$i). apply co_sizeof_pos. omega. Qed. (** The size of a type is an integral multiple of its alignment, @@ -380,8 +380,8 @@ Proof. - apply Z.divide_refl. - apply Z.divide_mul_l; auto. - apply Z.divide_refl. -- destruct (env!i). apply co_sizeof_alignof. apply Z.divide_0_r. -- destruct (env!i). apply co_sizeof_alignof. apply Z.divide_0_r. +- destruct (env$i). apply co_sizeof_alignof. apply Z.divide_0_r. +- destruct (env$i). apply co_sizeof_alignof. apply Z.divide_0_r. Qed. (** ** Size and alignment for composite definitions *) @@ -618,7 +618,7 @@ Fixpoint alignof_blockcopy (env: composite_env) (t: type) : Z := | Tarray t' _ _ => alignof_blockcopy env t' | Tfunction _ _ _ => 1 | Tstruct id _ | Tunion id _ => - match env!id with + match env$id with | Some co => Z.min 8 (co_alignof co) | None => 1 end @@ -646,8 +646,8 @@ Proof. destruct Archi.ptr64; auto. apply IHty. auto. - destruct (env!i); auto. - destruct (env!i); auto. + destruct (env$i); auto. + destruct (env$i); auto. Qed. Lemma alignof_blockcopy_pos: @@ -681,8 +681,8 @@ Proof. apply Z.divide_refl. apply Z.divide_mul_l. auto. apply Z.divide_refl. - destruct (env!i). apply X. apply Z.divide_0_r. - destruct (env!i). apply X. apply Z.divide_0_r. + destruct (env$i). apply X. apply Z.divide_0_r. + destruct (env$i). apply X. apply Z.divide_0_r. Qed. (** Type ranks *) @@ -697,7 +697,7 @@ Fixpoint rank_type (ce: composite_env) (t: type) : nat := match t with | Tarray t' _ _ => S (rank_type ce t') | Tstruct id _ | Tunion id _ => - match ce!id with + match ce$id with | None => O | Some co => S (co_rank co) end @@ -800,7 +800,7 @@ Qed. Program Definition composite_of_def (env: composite_env) (id: ident) (su: struct_or_union) (m: members) (a: attr) : res composite := - match env!id, complete_members env m return _ with + match env$id, complete_members env m return _ with | Some _, _ => Error (MSG "Multiple definitions of struct or union " :: CTX id :: nil) | None, false => @@ -841,11 +841,11 @@ Fixpoint add_composite_definitions (env: composite_env) (defs: list composite_de | nil => OK env | Composite id su m a :: defs => do co <- composite_of_def env id su m a; - add_composite_definitions (PTree.set id co env) defs + add_composite_definitions (ATree.set id co env) defs end. Definition build_composite_env (defs: list composite_definition) := - add_composite_definitions (PTree.empty _) defs. + add_composite_definitions (ATree.empty _) defs. (** Stability properties for alignments, sizes, and ranks. If the type is complete in a composite environment [env], its size, alignment, and rank @@ -854,15 +854,15 @@ Definition build_composite_env (defs: list composite_definition) := Section STABILITY. Variables env env': composite_env. -Hypothesis extends: forall id co, env!id = Some co -> env'!id = Some co. +Hypothesis extends: forall id co, env$id = Some co -> env'$id = Some co. Lemma alignof_stable: forall t, complete_type env t = true -> alignof env' t = alignof env t. Proof. induction t; simpl; intros; f_equal; auto. - destruct (env!i) as [co|] eqn:E; try discriminate. + destruct (env$i) as [co|] eqn:E; try discriminate. erewrite extends by eauto. auto. - destruct (env!i) as [co|] eqn:E; try discriminate. + destruct (env$i) as [co|] eqn:E; try discriminate. erewrite extends by eauto. auto. Qed. @@ -871,9 +871,9 @@ Lemma sizeof_stable: Proof. induction t; simpl; intros; auto. rewrite IHt by auto. auto. - destruct (env!i) as [co|] eqn:E; try discriminate. + destruct (env$i) as [co|] eqn:E; try discriminate. erewrite extends by eauto. auto. - destruct (env!i) as [co|] eqn:E; try discriminate. + destruct (env$i) as [co|] eqn:E; try discriminate. erewrite extends by eauto. auto. Qed. @@ -881,9 +881,9 @@ Lemma complete_type_stable: forall t, complete_type env t = true -> complete_type env' t = true. Proof. induction t; simpl; intros; auto. - destruct (env!i) as [co|] eqn:E; try discriminate. + destruct (env$i) as [co|] eqn:E; try discriminate. erewrite extends by eauto. auto. - destruct (env!i) as [co|] eqn:E; try discriminate. + destruct (env$i) as [co|] eqn:E; try discriminate. erewrite extends by eauto. auto. Qed. @@ -891,9 +891,9 @@ Lemma rank_type_stable: forall t, complete_type env t = true -> rank_type env' t = rank_type env t. Proof. induction t; simpl; intros; auto. - destruct (env!i) as [co|] eqn:E; try discriminate. + destruct (env$i) as [co|] eqn:E; try discriminate. erewrite extends by eauto. auto. - destruct (env!i) as [co|] eqn:E; try discriminate. + destruct (env$i) as [co|] eqn:E; try discriminate. erewrite extends by eauto. auto. Qed. @@ -951,12 +951,12 @@ End STABILITY. Lemma add_composite_definitions_incr: forall id co defs env1 env2, add_composite_definitions env1 defs = OK env2 -> - env1!id = Some co -> env2!id = Some co. + env1$id = Some co -> env2$id = Some co. Proof. induction defs; simpl; intros. - inv H; auto. - destruct a; monadInv H. - eapply IHdefs; eauto. rewrite PTree.gso; auto. + eapply IHdefs; eauto. rewrite ATree.gso; auto. red; intros; subst id0. unfold composite_of_def in EQ. rewrite H0 in EQ; discriminate. Qed. @@ -976,11 +976,11 @@ Record composite_consistent (env: composite_env) (co: composite) : Prop := { }. Definition composite_env_consistent (env: composite_env) : Prop := - forall id co, env!id = Some co -> composite_consistent env co. + forall id co, env$id = Some co -> composite_consistent env co. Lemma composite_consistent_stable: forall (env env': composite_env) - (EXTENDS: forall id co, env!id = Some co -> env'!id = Some co) + (EXTENDS: forall id co, env$id = Some co -> env'$id = Some co) co, composite_consistent env co -> composite_consistent env' co. Proof. @@ -997,7 +997,7 @@ Lemma composite_of_def_consistent: composite_consistent env co. Proof. unfold composite_of_def; intros. - destruct (env!id); try discriminate. destruct (complete_members env m) eqn:C; inv H. + destruct (env$id); try discriminate. destruct (complete_members env m) eqn:C; inv H. constructor; auto. Qed. @@ -1008,18 +1008,18 @@ Proof. add_composite_definitions env0 defs = OK env -> composite_env_consistent env0 -> composite_env_consistent env). - intros. eapply H; eauto. red; intros. rewrite PTree.gempty in H1; discriminate. + intros. eapply H; eauto. red; intros. rewrite ATree.gempty in H1; discriminate. induction defs as [|d1 defs]; simpl; intros. - inv H; auto. - destruct d1; monadInv H. eapply IHdefs; eauto. - set (env1 := PTree.set id x env0) in *. - assert (env0!id = None). - { unfold composite_of_def in EQ. destruct (env0!id). discriminate. auto. } - assert (forall id1 co1, env0!id1 = Some co1 -> env1!id1 = Some co1). - { intros. unfold env1. rewrite PTree.gso; auto. congruence. } + set (env1 := ATree.set id x env0) in *. + assert (env0$id = None). + { unfold composite_of_def in EQ. destruct (env0$id). discriminate. auto. } + assert (forall id1 co1, env0$id1 = Some co1 -> env1$id1 = Some co1). + { intros. unfold env1. rewrite ATree.gso; auto. congruence. } red; intros. apply composite_consistent_stable with env0; auto. - unfold env1 in H2; rewrite PTree.gsspec in H2; destruct (peq id0 id). + unfold env1 in H2; rewrite ATree.gsspec in H2; destruct (ATree.elt_eq id0 id). + subst id0. inversion H2; clear H2. subst co. eapply composite_of_def_consistent; eauto. + eapply H0; eauto. @@ -1031,42 +1031,42 @@ Theorem build_composite_env_charact: forall id su m a defs env, build_composite_env defs = OK env -> In (Composite id su m a) defs -> - exists co, env!id = Some co /\ co_members co = m /\ co_attr co = a /\ co_su co = su. + exists co, env$id = Some co /\ co_members co = m /\ co_attr co = a /\ co_su co = su. Proof. - intros until defs. unfold build_composite_env. generalize (PTree.empty composite) as env0. + intros until defs. unfold build_composite_env. generalize (ATree.empty composite) as env0. revert defs. induction defs as [|d1 defs]; simpl; intros. - contradiction. - destruct d1; monadInv H. destruct H0; [idtac|eapply IHdefs;eauto]. inv H. unfold composite_of_def in EQ. - destruct (env0!id) eqn:E; try discriminate. + destruct (env0$id) eqn:E; try discriminate. destruct (complete_members env0 m) eqn:C; simplify_eq EQ. clear EQ; intros EQ. exists x. - split. eapply add_composite_definitions_incr; eauto. apply PTree.gss. + split. eapply add_composite_definitions_incr; eauto. apply ATree.gss. subst x; auto. Qed. Theorem build_composite_env_domain: forall env defs id co, build_composite_env defs = OK env -> - env!id = Some co -> + env$id = Some co -> In (Composite id (co_su co) (co_members co) (co_attr co)) defs. Proof. intros env0 defs0 id co. assert (REC: forall l env env', add_composite_definitions env l = OK env' -> - env'!id = Some co -> - env!id = Some co \/ In (Composite id (co_su co) (co_members co) (co_attr co)) l). + env'$id = Some co -> + env$id = Some co \/ In (Composite id (co_su co) (co_members co) (co_attr co)) l). { induction l; simpl; intros. - inv H; auto. - destruct a; monadInv H. exploit IHl; eauto. - unfold composite_of_def in EQ. destruct (env!id0) eqn:E; try discriminate. + unfold composite_of_def in EQ. destruct (env$id0) eqn:E; try discriminate. destruct (complete_members env m) eqn:C; simplify_eq EQ. clear EQ; intros EQ. - rewrite PTree.gsspec. intros [A|A]; auto. - destruct (peq id id0); auto. + rewrite ATree.gsspec. intros [A|A]; auto. + destruct (_ id id0); auto. inv A. rewrite <- H0; auto. } - intros. exploit REC; eauto. rewrite PTree.gempty. intuition congruence. + intros. exploit REC; eauto. rewrite ATree.gempty. intuition congruence. Qed. (** As a corollay, in a consistent environment, the rank of a composite type @@ -1083,7 +1083,7 @@ Qed. Lemma rank_struct_member: forall ce id a co id1 t1, composite_env_consistent ce -> - ce!id = Some co -> + ce$id = Some co -> In (id1, t1) (co_members co) -> (rank_type ce t1 < rank_type ce (Tstruct id a))%nat. Proof. @@ -1096,7 +1096,7 @@ Qed. Lemma rank_union_member: forall ce id a co id1 t1, composite_env_consistent ce -> - ce!id = Some co -> + ce$id = Some co -> In (id1, t1) (co_members co) -> (rank_type ce t1 < rank_type ce (Tunion id a))%nat. Proof. @@ -1265,7 +1265,7 @@ Qed. Lemma composite_of_def_eq: forall env id co, composite_consistent env co -> - env!id = None -> + env$id = None -> composite_of_def env id (co_su co) (co_members co) (co_attr co) = OK co. Proof. intros. destruct H as [A B C D]. unfold composite_of_def. rewrite H0, A. @@ -1286,15 +1286,15 @@ Qed. Lemma composite_of_def_stable: forall (env env': composite_env) - (EXTENDS: forall id co, env!id = Some co -> env'!id = Some co) + (EXTENDS: forall id co, env$id = Some co -> env'$id = Some co) id su m a co, - env'!id = None -> + env'$id = None -> composite_of_def env id su m a = OK co -> composite_of_def env' id su m a = OK co. Proof. intros. unfold composite_of_def in H0. - destruct (env!id) eqn:E; try discriminate. + destruct (env$id) eqn:E; try discriminate. destruct (complete_members env m) eqn:CM; try discriminate. transitivity (composite_of_def env' id (co_su co) (co_members co) (co_attr co)). inv H0; auto. @@ -1308,26 +1308,26 @@ Lemma link_add_composite_definitions: build_composite_env l0 = OK env0 -> forall l env1 env1' env2, add_composite_definitions env1 l = OK env1' -> - (forall id co, env1!id = Some co -> env2!id = Some co) -> - (forall id co, env0!id = Some co -> env2!id = Some co) -> - (forall id, env2!id = if In_dec ident_eq id (map name_composite_def l0) then env0!id else env1!id) -> + (forall id co, env1$id = Some co -> env2$id = Some co) -> + (forall id co, env0$id = Some co -> env2$id = Some co) -> + (forall id, env2$id = if In_dec ident_eq id (map name_composite_def l0) then env0$id else env1$id) -> ((forall cd1 cd2, In cd1 l0 -> In cd2 l -> name_composite_def cd2 = name_composite_def cd1 -> cd2 = cd1)) -> { env2' | add_composite_definitions env2 (filter_redefs l0 l) = OK env2' - /\ (forall id co, env1'!id = Some co -> env2'!id = Some co) - /\ (forall id co, env0!id = Some co -> env2'!id = Some co) }. + /\ (forall id co, env1'$id = Some co -> env2'$id = Some co) + /\ (forall id co, env0$id = Some co -> env2'$id = Some co) }. Proof. induction l; simpl; intros until env2; intros ACD AGREE1 AGREE0 AGREE2 UNIQUE. - inv ACD. exists env2; auto. - destruct a. destruct (composite_of_def env1 id su m a) as [x|e] eqn:EQ; try discriminate. simpl in ACD. generalize EQ. unfold composite_of_def at 1. - destruct (env1!id) eqn:E1; try congruence. + destruct (env1$id) eqn:E1; try congruence. destruct (complete_members env1 m) eqn:CM1; try congruence. intros EQ1. simpl. destruct (in_dec ident_eq id (map name_composite_def l0)); simpl. + eapply IHl; eauto. -* intros. rewrite PTree.gsspec in H0. destruct (peq id0 id); auto. +* intros. rewrite ATree.gsspec in H0. destruct (ATree.elt_eq id0 id); auto. inv H0. exploit list_in_map_inv; eauto. intros ([id' su' m' a'] & P & Q). assert (X: Composite id su m a = Composite id' su' m' a'). @@ -1345,16 +1345,16 @@ Proof. inversion EQ1; auto. } subst co'. apply AGREE0; auto. * intros. rewrite AGREE2. destruct (in_dec ident_eq id0 (map name_composite_def l0)); auto. - rewrite PTree.gsspec. destruct (peq id0 id); auto. subst id0. contradiction. -+ assert (E2: env2!id = None). + rewrite ATree.gsspec. destruct (ATree.elt_eq id0 id); auto. subst id0. contradiction. ++ assert (E2: env2$id = None). { rewrite AGREE2. rewrite pred_dec_false by auto. auto. } assert (E3: composite_of_def env2 id su m a = OK x). { eapply composite_of_def_stable. eexact AGREE1. eauto. eauto. } rewrite E3. simpl. eapply IHl; eauto. -* intros until co; rewrite ! PTree.gsspec. destruct (peq id0 id); auto. -* intros until co; rewrite ! PTree.gsspec. intros. destruct (peq id0 id); auto. +* intros until co; rewrite ! ATree.gsspec. destruct (ATree.elt_eq id0 id); auto. +* intros until co; rewrite ! ATree.gsspec. intros. destruct (ATree.elt_eq id0 id); auto. subst id0. apply AGREE0 in H0. congruence. -* intros. rewrite ! PTree.gsspec. destruct (peq id0 id); auto. subst id0. +* intros. rewrite ! ATree.gsspec. destruct (ATree.elt_eq id0 id); auto. subst id0. rewrite pred_dec_false by auto. auto. Qed. @@ -1365,17 +1365,17 @@ Theorem link_build_composite_env: link l1 l2 = Some l -> { env | build_composite_env l = OK env - /\ (forall id co, env1!id = Some co -> env!id = Some co) - /\ (forall id co, env2!id = Some co -> env!id = Some co) }. + /\ (forall id co, env1$id = Some co -> env$id = Some co) + /\ (forall id co, env2$id = Some co -> env$id = Some co) }. Proof. intros. edestruct link_composite_def_inv as (A & B & C); eauto. edestruct link_add_composite_definitions as (env & P & Q & R). eexact H. eexact H0. - instantiate (1 := env1). intros. rewrite PTree.gempty in H2; discriminate. + instantiate (1 := env1). intros. rewrite ATree.gempty in H2; discriminate. auto. intros. destruct (in_dec ident_eq id (map name_composite_def l1)); auto. - rewrite PTree.gempty. destruct (env1!id) eqn:E1; auto. + rewrite ATree.gempty. destruct (env1$id) eqn:E1; auto. exploit build_composite_env_domain. eexact H. eauto. intros. apply (in_map name_composite_def) in H2. elim n; auto. auto. @@ -1463,7 +1463,7 @@ Definition link_program {F:Type} (p1 p2: program F): option (program F) := Definition linkorder_program {F: Type} (p1 p2: program F) : Prop := linkorder (program_of_program p1) (program_of_program p2) - /\ (forall id co, p1.(prog_comp_env)!id = Some co -> p2.(prog_comp_env)!id = Some co). + /\ (forall id co, p1.(prog_comp_env)$id = Some co -> p2.(prog_comp_env)$id = Some co). Instance Linker_program (F: Type): Linker (program F) := { link := link_program; diff --git a/cfrontend/Ctyping.v b/cfrontend/Ctyping.v index ba1d34fbdd..6120269d9c 100644 --- a/cfrontend/Ctyping.v +++ b/cfrontend/Ctyping.v @@ -258,7 +258,7 @@ Definition type_conditional (ty1 ty2: type) : res type := (** Type environments map identifiers to their types. *) -Definition typenv := PTree.t type. +Definition typenv := ATree.t type. Definition wt_cast (from to: type) : Prop := classify_cast from to <> cast_case_default. @@ -408,7 +408,7 @@ with wt_lvalue : expr -> Prop := | wt_Eloc: forall b ofs ty, wt_lvalue (Eloc b ofs ty) | wt_Evar: forall x ty, - e!x = Some ty -> + e$x = Some ty -> wt_lvalue (Evar x ty) | wt_Ederef: forall r ty, wt_rvalue r -> @@ -417,7 +417,7 @@ with wt_lvalue : expr -> Prop := | wt_Efield: forall r f id a co ty, wt_rvalue r -> typeof r = Tstruct id a \/ typeof r = Tunion id a -> - ce!id = Some co -> + ce$id = Some co -> type_of_member a f co.(co_members) = OK ty -> wt_lvalue (Efield r f ty) @@ -502,7 +502,7 @@ End WT_EXPR_STMT. Fixpoint bind_vars (e: typenv) (l: list (ident * type)) : typenv := match l with | nil => e - | (id, ty) :: l => bind_vars (PTree.set id ty e) l + | (id, ty) :: l => bind_vars (ATree.set id ty e) l end. Inductive wt_function (ce: composite_env) (e: typenv) : function -> Prop := @@ -513,13 +513,13 @@ Inductive wt_function (ce: composite_env) (e: typenv) : function -> Prop := Fixpoint bind_globdef (e: typenv) (l: list (ident * globdef fundef type)) : typenv := match l with | nil => e - | (id, Gfun fd) :: l => bind_globdef (PTree.set id (type_of_fundef fd) e) l - | (id, Gvar v) :: l => bind_globdef (PTree.set id v.(gvar_info) e) l + | (id, Gfun fd) :: l => bind_globdef (ATree.set id (type_of_fundef fd) e) l + | (id, Gvar v) :: l => bind_globdef (ATree.set id v.(gvar_info) e) l end. Inductive wt_program : program -> Prop := | wt_program_intro: forall p, - let e := bind_globdef (PTree.empty _) p.(prog_defs) in + let e := bind_globdef (ATree.empty _) p.(prog_defs) in (forall id f, In (id, Gfun (Internal f)) p.(prog_defs) -> wt_function p.(prog_comp_env) e f) -> wt_program p. @@ -607,7 +607,7 @@ Fixpoint check_rvals (el: exprlist) : res unit := type annotation. *) Definition evar (e: typenv) (x: ident) : res expr := - match e!x with + match e$x with | Some ty => OK (Evar x ty) | None => Error (MSG "unbound variable " :: CTX x :: nil) end. @@ -621,7 +621,7 @@ Definition efield (ce: composite_env) (r: expr) (f: ident) : res expr := do x1 <- check_rval r; match typeof r with | Tstruct id a | Tunion id a => - match ce!id with + match ce$id with | None => Error (MSG "unbound composite " :: CTX id :: nil) | Some co => do ty <- type_of_member a f co.(co_members); @@ -909,7 +909,7 @@ Definition retype_fundef (ce: composite_env) (e: typenv) (fd: fundef) : res fund end. Definition typecheck_program (p: program) : res program := - let e := bind_globdef (PTree.empty _) p.(prog_defs) in + let e := bind_globdef (ATree.empty _) p.(prog_defs) in let ce := p.(prog_comp_env) in do tp <- transform_partial_program (retype_fundef ce e) p; OK {| prog_defs := tp.(AST.prog_defs); @@ -1062,7 +1062,7 @@ Hint Extern 1 (wt_expr _ _ _) => (unfold wt_expr; simpl): ty. Lemma evar_sound: forall x a, evar e x = OK a -> wt_expr ce e a. Proof. - unfold evar; intros. destruct (e!x) as [ty|] eqn:E; inv H. eauto with ty. + unfold evar; intros. destruct (e$x) as [ty|] eqn:E; inv H. eauto with ty. Qed. Lemma ederef_sound: @@ -1076,7 +1076,7 @@ Lemma efield_sound: Proof. intros. monadInv H. destruct (typeof r) eqn:TR; try discriminate; - destruct (ce!i) as [co|] eqn:CE; monadInv EQ0; eauto with ty. + destruct (ce$i) as [co|] eqn:CE; monadInv EQ0; eauto with ty. Qed. Lemma econst_int_sound: @@ -1349,13 +1349,13 @@ Proof. rename x into tp. constructor; simpl. set (ce := prog_comp_env p) in *. - set (e := bind_globdef (PTree.empty type) (prog_defs p)) in *. - set (e' := bind_globdef (PTree.empty type) (AST.prog_defs tp)) in *. + set (e := bind_globdef (ATree.empty type) (prog_defs p)) in *. + set (e' := bind_globdef (ATree.empty type) (AST.prog_defs tp)) in *. assert (M: match_program (fun ctx f tf => retype_fundef ce e f = OK tf) eq p tp) by (eapply match_transform_partial_program; eauto). destruct M as (MATCH & _). simpl in MATCH. assert (ENVS: e' = e). - { unfold e, e'. revert MATCH; generalize (prog_defs p) (AST.prog_defs tp) (PTree.empty type). + { unfold e, e'. revert MATCH; generalize (prog_defs p) (AST.prog_defs tp) (ATree.empty type). induction l as [ | [id gd] l ]; intros l' t M; inv M. auto. destruct b1 as [id' gd']; destruct H1; simpl in *. inv H0; simpl. @@ -1852,7 +1852,7 @@ Section PRESERVATION. Variable prog: program. Hypothesis WTPROG: wt_program prog. Let ge := globalenv prog. -Let gtenv := bind_globdef (PTree.empty _) prog.(prog_defs). +Let gtenv := bind_globdef (ATree.empty _) prog.(prog_defs). Hypothesis WT_EXTERNAL: forall id ef args res cc vargs m t vres m', diff --git a/cfrontend/Initializers.v b/cfrontend/Initializers.v index 77d6cfea85..27b384d44d 100644 --- a/cfrontend/Initializers.v +++ b/cfrontend/Initializers.v @@ -44,7 +44,7 @@ Definition do_cast (v: val) (t1 t2: type) : res val := end. Definition lookup_composite (ce: composite_env) (id: ident) : res composite := - match ce!id with + match ce$id with | Some co => OK co | None => Error (MSG "Undefined struct or union " :: CTX id :: nil) end. @@ -110,7 +110,7 @@ Fixpoint constval (ce: composite_env) (a: expr) : res val := | Ecomma r1 r2 ty => do v1 <- constval ce r1; constval ce r2 | Evar x ty => - OK(Vptr x Ptrofs.zero) + OK(Vptr (Ident.to_positive x) Ptrofs.zero) | Ederef r ty => constval ce r | Efield l f ty => @@ -163,9 +163,9 @@ Definition transl_init_single (ce: composite_env) (ty: type) (a: expr) : res ini | Vlong n, Tpointer _ _ => assertion (Archi.ptr64); OK(Init_int64 n) | Vsingle f, Tfloat F32 _ => OK(Init_float32 f) | Vfloat f, Tfloat F64 _ => OK(Init_float64 f) - | Vptr id ofs, Tint I32 sg _ => assertion (negb Archi.ptr64); OK(Init_addrof id ofs) - | Vptr id ofs, Tlong _ _ => assertion (Archi.ptr64); OK(Init_addrof id ofs) - | Vptr id ofs, Tpointer _ _ => OK(Init_addrof id ofs) + | Vptr id ofs, Tint I32 sg _ => assertion (negb Archi.ptr64); OK(Init_addrof (Ident.of_positive id) ofs) + | Vptr id ofs, Tlong _ _ => assertion (Archi.ptr64); OK(Init_addrof (Ident.of_positive id) ofs) + | Vptr id ofs, Tpointer _ _ => OK(Init_addrof (Ident.of_positive id) ofs) | Vundef, _ => Error(msg "undefined operation in initializer") | _, _ => Error (msg "type mismatch in initializer") end. diff --git a/cfrontend/Initializersproof.v b/cfrontend/Initializersproof.v index 272b929f1a..f2089226dc 100644 --- a/cfrontend/Initializersproof.v +++ b/cfrontend/Initializersproof.v @@ -69,10 +69,10 @@ Inductive eval_simple_lvalue: expr -> block -> ptrofs -> Prop := | esl_loc: forall b ofs ty, eval_simple_lvalue (Eloc b ofs ty) b ofs | esl_var_local: forall x ty b, - e!x = Some(b, ty) -> + e$x = Some(b, ty) -> eval_simple_lvalue (Evar x ty) b Ptrofs.zero | esl_var_global: forall x ty b, - e!x = None -> + e$x = None -> Genv.find_symbol ge x = Some b -> eval_simple_lvalue (Evar x ty) b Ptrofs.zero | esl_deref: forall r ty b ofs, @@ -80,7 +80,7 @@ Inductive eval_simple_lvalue: expr -> block -> ptrofs -> Prop := eval_simple_lvalue (Ederef r ty) b ofs | esl_field_struct: forall r f ty b ofs id co a delta, eval_simple_rvalue r (Vptr b ofs) -> - typeof r = Tstruct id a -> ge.(genv_cenv)!id = Some co -> field_offset ge f (co_members co) = OK delta -> + typeof r = Tstruct id a -> ge.(genv_cenv)$id = Some co -> field_offset ge f (co_members co) = OK delta -> eval_simple_lvalue (Efield r f ty) b (Ptrofs.add ofs (Ptrofs.repr delta)) | esl_field_union: forall r f ty b ofs id a, eval_simple_rvalue r (Vptr b ofs) -> @@ -323,7 +323,7 @@ Qed. [Vptr b ofs] where [Genv.find_symbol ge id = Some b]. *) Definition inj (b: block) := - match Genv.find_symbol ge b with + match Genv.find_symbol ge (Ident.of_positive b) with | Some b' => Some (b', 0) | None => None end. @@ -438,9 +438,9 @@ Proof. (* lvalue *) induction 1; intros v' CV; simpl in CV; try (monadInv CV). (* var local *) - unfold empty_env in H. rewrite PTree.gempty in H. congruence. + unfold empty_env in H. rewrite ATree.gempty in H. congruence. (* var_global *) - econstructor. unfold inj. rewrite H0. eauto. auto. + econstructor. unfold inj. rewrite Ident.of_to_positive, H0. eauto. auto. (* deref *) eauto. (* field struct *) @@ -627,13 +627,13 @@ Proof. destruct f1; inv EQ0; simpl in H2; inv H2; assumption. - (* pointer *) unfold inj in H. - assert (data = Init_addrof b1 ofs1 /\ chunk = Mptr). + assert (data = Init_addrof (Ident.of_positive b1) ofs1 /\ chunk = Mptr). { remember Archi.ptr64 as ptr64. destruct ty; inversion EQ0. destruct i; inv H5. unfold Mptr. destruct Archi.ptr64; inv H6; inv H2; auto. subst ptr64. unfold Mptr. destruct Archi.ptr64; inv H5; inv H2; auto. inv H2. auto. } - destruct H4; subst. destruct (Genv.find_symbol ge b1); inv H. + destruct H4; subst. destruct Genv.find_symbol; inv H. rewrite Ptrofs.add_zero in H3. auto. - (* undef *) discriminate. @@ -716,13 +716,13 @@ Local Opaque sizeof. + Local Transparent sizeof. simpl. eapply tr_init_array_size; eauto. + replace (idlsize d) with (idlsize d + 0) by omega. eapply tr_init_struct_size; eauto. simpl. - unfold lookup_composite in H. destruct (ge.(genv_cenv)!id) as [co'|] eqn:?; inv H. + unfold lookup_composite in H. destruct (ge.(genv_cenv)$id) as [co'|] eqn:?; inv H. erewrite co_consistent_sizeof by (eapply ce_consistent; eauto). unfold sizeof_composite. rewrite H0. apply align_le. destruct (co_alignof_two_p co) as [n EQ]. rewrite EQ. apply two_power_nat_pos. + rewrite idlsize_app, padding_size. exploit tr_init_size; eauto. intros EQ; rewrite EQ. omega. - simpl. unfold lookup_composite in H. destruct (ge.(genv_cenv)!id) as [co'|] eqn:?; inv H. + simpl. unfold lookup_composite in H. destruct (ge.(genv_cenv)$id) as [co'|] eqn:?; inv H. apply Z.le_trans with (sizeof_union ge (co_members co)). eapply union_field_size; eauto. erewrite co_consistent_sizeof by (eapply ce_consistent; eauto). @@ -771,11 +771,11 @@ Inductive exec_init: mem -> block -> Z -> type -> initializer -> mem -> Prop := exec_init_array m b ofs ty sz il m' -> exec_init m b ofs (Tarray ty sz a) (Init_array il) m' | exec_init_struct: forall m b ofs id a il co m', - ge.(genv_cenv)!id = Some co -> co_su co = Struct -> + ge.(genv_cenv)$id = Some co -> co_su co = Struct -> exec_init_list m b ofs (fields_of_struct (co_members co) 0) il m' -> exec_init m b ofs (Tstruct id a) (Init_struct il) m' | exec_init_union: forall m b ofs id a f i ty co m', - ge.(genv_cenv)!id = Some co -> co_su co = Union -> + ge.(genv_cenv)$id = Some co -> co_su co = Union -> field_type f (co_members co) = OK ty -> exec_init m b ofs ty i m' -> exec_init m b ofs (Tunion id a) (Init_union f i) m' diff --git a/cfrontend/PrintClight.ml b/cfrontend/PrintClight.ml index ca378c1126..d2f19db236 100644 --- a/cfrontend/PrintClight.ml +++ b/cfrontend/PrintClight.ml @@ -23,10 +23,6 @@ open Cop open PrintCsyntax open Clight -(* Naming temporaries *) - -let temp_name (id: AST.ident) = "$" ^ Z.to_string (Z.Zpos id) - (* Declarator (identifier + type) -- reuse from PrintCsyntax *) (* Precedences and associativity (H&S section 7.2) *) @@ -69,7 +65,7 @@ let rec expr p (prec, e) = | Evar(id, _) -> fprintf p "%s" (extern_atom id) | Etempvar(id, _) -> - fprintf p "%s" (temp_name id) + fprintf p "%s" (extern_atom id) | Ederef(a1, _) -> fprintf p "*%a" expr (prec', a1) | Efield(a1, f, _) -> @@ -123,14 +119,14 @@ let rec print_stmt p s = | Sassign(e1, e2) -> fprintf p "@[%a =@ %a;@]" print_expr e1 print_expr e2 | Sset(id, e2) -> - fprintf p "@[%s =@ %a;@]" (temp_name id) print_expr e2 + fprintf p "@[%s =@ %a;@]" (extern_atom id) print_expr e2 | Scall(None, e1, el) -> fprintf p "@[%a@,(@[%a@]);@]" expr (15, e1) print_expr_list (true, el) | Scall(Some id, e1, el) -> fprintf p "@[%s =@ %a@,(@[%a@]);@]" - (temp_name id) + (extern_atom id) expr (15, e1) print_expr_list (true, el) | Sbuiltin(None, ef, tyargs, el) -> @@ -139,7 +135,7 @@ let rec print_stmt p s = print_expr_list (true, el) | Sbuiltin(Some id, ef, tyargs, el) -> fprintf p "@[%s =@ builtin %s@,(@[%a@]);@]" - (temp_name id) + (extern_atom id) (name_of_external ef) print_expr_list (true, el) | Ssequence(Sskip, s2) -> @@ -210,7 +206,7 @@ and print_stmt_for p s = | Sassign(e1, e2) -> fprintf p "%a = %a" print_expr e1 print_expr e2 | Sset(id, e2) -> - fprintf p "%s = %a" (temp_name id) print_expr e2 + fprintf p "%s = %a" (extern_atom id) print_expr e2 | Ssequence(Sskip, s2) -> print_stmt_for p s2 | Ssequence(s1, s2) -> @@ -221,7 +217,7 @@ and print_stmt_for p s = print_expr_list (true, el) | Scall(Some id, e1, el) -> fprintf p "@[%s =@ %a@,(@[%a@])@]" - (temp_name id) + (extern_atom id) expr (15, e1) print_expr_list (true, el) | Sbuiltin(None, ef, tyargs, el) -> @@ -230,7 +226,7 @@ and print_stmt_for p s = print_expr_list (true, el) | Sbuiltin(Some id, ef, tyargs, el) -> fprintf p "@[%s =@ builtin %s@,(@[%a@]);@]" - (temp_name id) + (extern_atom id) (name_of_external ef) print_expr_list (true, el) | _ -> @@ -248,7 +244,7 @@ let print_function p id f = f.fn_vars; List.iter (fun (id, ty) -> - fprintf p "register %s;@ " (name_cdecl (temp_name id) ty)) + fprintf p "register %s;@ " (name_cdecl (extern_atom id) ty)) f.fn_temps; print_stmt p f.fn_body; fprintf p "@;<0 -2>}@]@ @ " diff --git a/cfrontend/SimplExpr.v b/cfrontend/SimplExpr.v index 45b686f3d3..f3a36f8672 100644 --- a/cfrontend/SimplExpr.v +++ b/cfrontend/SimplExpr.v @@ -13,6 +13,7 @@ (** Translation from Compcert C to Clight. Side effects are pulled out of Compcert C expressions. *) +Require Import String. Require Import Coqlib. Require Import Errors. Require Import Integers. @@ -25,12 +26,10 @@ Require Import Cop. Require Import Csyntax. Require Import Clight. -Local Open Scope string_scope. - (** State and error monad for generating fresh identifiers. *) Record generator : Type := mkgenerator { - gen_next: ident; + gen_next: positive; gen_trail: list (ident * type) }. @@ -72,15 +71,58 @@ Notation "'do' ( X , Y ) <- A ; B" := (bind2 A (fun X Y => B)) Local Open Scope gensym_monad_scope. -Parameter first_unused_ident: unit -> ident. - Definition initial_generator (x: unit) : generator := - mkgenerator (first_unused_ident x) nil. + mkgenerator 1%positive nil. + +(** We need to generate symbols of the form "$n", where n is a decimal + representation for the current value of genv_next. Because such + symbols are not legal C identifiers, they will not conflict with the + ones we got from the parser. *) + +Lemma N_size_nat_le m n: + N.le m n -> le (N.size_nat m) (N.size_nat n). +Proof. + intros H. + apply N.right_induction with (z := m) (n := n); eauto. + - intros x y []. tauto. + - intros. + transitivity (N.size_nat n0); eauto. + clear. + destruct n0; simpl. { apply Nat.le_0_1. } + induction p; simpl; eauto using le_n_S. +Qed. + +Definition digit_of (n: N) := + String (Ascii.ascii_of_N (48 + n)) EmptyString. + +Function digits_of (n: N) {measure N.size_nat} := + if N.eqb n 0%N then + EmptyString + else + let (q, r) := N.div_eucl n 10%N in + String.append (digits_of q) (digit_of r). +Proof. + intros n Hnz q r Hqr. + apply N.eqb_neq in Hnz. destruct n as [|n]; try congruence. simpl in Hqr. + destruct q as [|q]. { destruct n; apply Nat.lt_0_succ. } + pose proof (N.pos_div_eucl_spec n 10) as Hspec. + rewrite Hqr in Hspec. rewrite Hspec. + apply lt_le_trans with (N.size_nat (N.pos q * 2)). + { + rewrite N.mul_comm. simpl. + apply Nat.lt_succ_diag_r. + } + apply N_size_nat_le. + xomega. +Defined. + +Definition string_of_resid n := + String "$" (digits_of (Npos n)). Definition gensym (ty: type): mon ident := fun (g: generator) => - Res (gen_next g) - (mkgenerator (Pos.succ (gen_next g)) ((gen_next g, ty) :: gen_trail g)) + Res #(string_of_resid (gen_next g)) + (mkgenerator (Pos.succ (gen_next g)) ((#(string_of_resid (gen_next g)), ty) :: gen_trail g)) (Ple_succ (gen_next g)). (** Construct a sequence from a list of statements. To facilitate the diff --git a/cfrontend/SimplExprproof.v b/cfrontend/SimplExprproof.v index ee1df409d5..062e379770 100644 --- a/cfrontend/SimplExprproof.v +++ b/cfrontend/SimplExprproof.v @@ -367,7 +367,7 @@ Lemma tr_expr_leftcontext_rec: /\ incl tmp' tmps /\ (forall le' e' sl3, tr_expr le' dst' e' sl3 a' tmp' -> - (forall id, ~In id tmp' -> le'!id = le!id) -> + (forall id, ~In id tmp' -> le'$id = le$id) -> Csyntax.typeof e' = Csyntax.typeof e -> tr_expr le' dst (C e') (sl3 ++ sl2) a tmps) ) /\ ( @@ -380,7 +380,7 @@ Lemma tr_expr_leftcontext_rec: /\ incl tmp' tmps /\ (forall le' e' sl3, tr_expr le' dst' e' sl3 a' tmp' -> - (forall id, ~In id tmp' -> le'!id = le!id) -> + (forall id, ~In id tmp' -> le'$id = le$id) -> Csyntax.typeof e' = Csyntax.typeof e -> tr_exprlist le' (C e') (sl3 ++ sl2) a tmps) ). @@ -400,8 +400,8 @@ Ltac NOTIN := Ltac UNCHANGED := match goal with - | [ H: (forall (id: ident), ~In id _ -> ?le' ! id = ?le ! id) |- - (forall (id: ident), In id _ -> ?le' ! id = ?le ! id) ] => + | [ H: (forall (id: ident), ~In id _ -> ?le' $ id = ?le $ id) |- + (forall (id: ident), In id _ -> ?le' $ id = ?le $ id) ] => intros; apply H; NOTIN end. @@ -725,7 +725,7 @@ Theorem tr_expr_leftcontext: /\ incl tmp' tmps /\ (forall le' r' sl3, tr_expr le' dst' r' sl3 a' tmp' -> - (forall id, ~In id tmp' -> le'!id = le!id) -> + (forall id, ~In id tmp' -> le'$id = le$id) -> Csyntax.typeof r' = Csyntax.typeof r -> tr_expr le' dst (C r') (sl3 ++ sl2) a tmps). Proof. @@ -744,7 +744,7 @@ Theorem tr_top_leftcontext: /\ incl tmp' tmps /\ (forall le' m' r' sl3, tr_expr le' dst' r' sl3 a' tmp' -> - (forall id, ~In id tmp' -> le'!id = le!id) -> + (forall id, ~In id tmp' -> le'$id = le$id) -> Csyntax.typeof r' = Csyntax.typeof r -> tr_top tge e le' m' dst (C r') (sl3 ++ sl2) a tmps). Proof. @@ -833,12 +833,12 @@ Lemma step_make_set: eval_lvalue tge e le m a b ofs -> typeof a = ty -> step1 tge (State f (make_set id a) k e le m) - t (State f Sskip k e (PTree.set id v le) m). + t (State f Sskip k e (ATree.set id v le) m). Proof. intros. exploit deref_loc_translated; eauto. rewrite <- H1. unfold make_set. destruct (chunk_for_volatile_type (typeof a)) as [chunk|]. (* volatile case *) - intros. change (PTree.set id v le) with (set_opttemp (Some id) v le). econstructor. + intros. change (ATree.set id v le) with (set_opttemp (Some id) v le). econstructor. econstructor. constructor. eauto. simpl. unfold sem_cast. simpl. eauto. constructor. simpl. econstructor; eauto. @@ -903,7 +903,7 @@ Lemma step_tr_rvalof: t (State f Sskip k e le' m) /\ eval_expr tge e le' m a' v /\ typeof a' = typeof a - /\ forall x, ~In x tmp -> le'!x = le!x. + /\ forall x, ~In x tmp -> le'$x = le$x. Proof. intros. inv H1. (* not volatile *) @@ -913,11 +913,11 @@ Proof. split. eapply eval_Elvalue; eauto. auto. (* volatile *) - intros. exists (PTree.set t0 v le); split. + intros. exists (ATree.set t0 v le); split. simpl. eapply star_two. econstructor. eapply step_make_set; eauto. traceEq. - split. constructor. apply PTree.gss. + split. constructor. apply ATree.gss. split. auto. - intros. apply PTree.gso. congruence. + intros. apply ATree.gso. congruence. Qed. (** Matching between continuations *) @@ -1204,7 +1204,7 @@ Lemma tr_find_label_expression: Proof. intros. inv H. assert (nolabel (makeseq sl)). apply makeseq_nolabel. - eapply tr_find_label_top with (e := empty_env) (le := PTree.empty val) (m := Mem.empty). + eapply tr_find_label_top with (e := empty_env) (le := ATree.empty val) (m := Mem.empty). eauto. apply H. Qed. @@ -1213,7 +1213,7 @@ Lemma tr_find_label_expr_stmt: Proof. intros. inv H. assert (nolabel (makeseq sl)). apply makeseq_nolabel. - eapply tr_find_label_top with (e := empty_env) (le := PTree.empty val) (m := Mem.empty). + eapply tr_find_label_top with (e := empty_env) (le := ATree.empty val) (m := Mem.empty). eauto. apply H. Qed. @@ -1226,7 +1226,7 @@ Proof. assert (nolabel (makeseq (sl ++ makeif a Sskip Sbreak :: nil))). apply makeseq_nolabel. apply nolabel_list_app. - eapply tr_find_label_top with (e := empty_env) (le := PTree.empty val) (m := Mem.empty). + eapply tr_find_label_top with (e := empty_env) (le := ATree.empty val) (m := Mem.empty). eauto. simpl; split; auto. apply makeif_nolabel. red; simpl; auto. red; simpl; auto. apply H. @@ -1419,7 +1419,7 @@ Lemma tr_val_gen: forall le dst v ty a tmp, typeof a = ty -> (forall tge e le' m, - (forall id, In id tmp -> le'!id = le!id) -> + (forall id, In id tmp -> le'$id = le$id) -> eval_expr tge e le' m a v) -> tr_expr le dst (Csyntax.Eval v ty) (final dst a) a tmp. Proof. @@ -1463,8 +1463,8 @@ Proof. econstructor; eauto. change (final dst' (Etempvar t0 (Csyntax.typeof l)) ++ sl2) with (nil ++ (final dst' (Etempvar t0 (Csyntax.typeof l)) ++ sl2)). apply S. apply tr_val_gen. auto. - intros. constructor. rewrite H5; auto. apply PTree.gss. - intros. apply PTree.gso. red; intros; subst; elim H5; auto. + intros. constructor. rewrite H5; auto. apply ATree.gss. + intros. apply ATree.gso. red; intros; subst; elim H5; auto. auto. (* seqand true *) exploit tr_top_leftcontext; eauto. clear H9. @@ -1516,8 +1516,8 @@ Proof. apply star_one. constructor. constructor. reflexivity. reflexivity. eapply match_exprstates; eauto. change sl2 with (nil ++ sl2). apply S. econstructor; eauto. - intros. constructor. rewrite H2. apply PTree.gss. auto. - intros. apply PTree.gso. congruence. + intros. constructor. rewrite H2. apply ATree.gss. auto. + intros. apply ATree.gso. congruence. auto. (* for effects *) exploit tr_simple_rvalue; eauto. intros [SL [TY EV]]. @@ -1552,8 +1552,8 @@ Proof. apply star_one. constructor. constructor. reflexivity. reflexivity. eapply match_exprstates; eauto. change sl2 with (nil ++ sl2). apply S. econstructor; eauto. - intros. constructor. rewrite H2. apply PTree.gss. auto. - intros. apply PTree.gso. congruence. + intros. constructor. rewrite H2. apply ATree.gss. auto. + intros. apply ATree.gso. congruence. auto. (* for effects *) exploit tr_simple_rvalue; eauto. intros [SL [TY EV]]. @@ -1696,8 +1696,8 @@ Proof. (* for value *) exploit tr_simple_rvalue; eauto. intros [SL2 [TY2 EV2]]. exploit tr_simple_lvalue. eauto. - eapply tr_expr_invariant with (le' := PTree.set t0 v' le). eauto. - intros. apply PTree.gso. intuition congruence. + eapply tr_expr_invariant with (le' := ATree.set t0 v' le). eauto. + intros. apply ATree.gso. intuition congruence. intros [SL1 [TY1 EV1]]. subst; simpl Kseqlist. econstructor; split. @@ -1705,12 +1705,12 @@ Proof. eapply star_left. constructor. econstructor. eauto. rewrite <- TY2; eauto. eapply star_left. constructor. apply star_one. eapply step_make_assign; eauto. - constructor. apply PTree.gss. simpl. eapply cast_idempotent; eauto. + constructor. apply ATree.gss. simpl. eapply cast_idempotent; eauto. reflexivity. reflexivity. traceEq. econstructor. auto. apply S. apply tr_val_gen. auto. intros. constructor. - rewrite H4; auto. apply PTree.gss. - intros. apply PTree.gso. intuition congruence. + rewrite H4; auto. apply ATree.gss. + intros. apply ATree.gso. intuition congruence. auto. auto. (* assignop *) exploit tr_top_leftcontext; eauto. clear H15. @@ -1740,8 +1740,8 @@ Proof. exploit tr_simple_rvalue. eauto. eapply tr_expr_invariant with (le := le) (le' := le'). eauto. intros. apply INV. NOTIN. simpl. intros [SL2 [TY2 EV2]]. exploit tr_simple_lvalue. eauto. - eapply tr_expr_invariant with (le := le) (le' := PTree.set t v4 le'). eauto. - intros. rewrite PTree.gso. apply INV. NOTIN. intuition congruence. + eapply tr_expr_invariant with (le := le) (le' := ATree.set t v4 le'). eauto. + intros. rewrite ATree.gso. apply INV. NOTIN. intuition congruence. intros [? [? EV1'']]. subst; simpl Kseqlist. econstructor; split. @@ -1752,12 +1752,12 @@ Proof. rewrite TY3; rewrite <- TY1; rewrite <- TY2; rewrite comp_env_preserved; eauto. eassumption. econstructor. eapply step_make_assign; eauto. - constructor. apply PTree.gss. simpl. eapply cast_idempotent; eauto. + constructor. apply ATree.gss. simpl. eapply cast_idempotent; eauto. reflexivity. traceEq. econstructor. auto. apply S. apply tr_val_gen. auto. intros. constructor. - rewrite H10; auto. apply PTree.gss. - intros. rewrite PTree.gso. apply INV. + rewrite H10; auto. apply ATree.gss. + intros. rewrite ATree.gso. apply INV. red; intros; elim H10; auto. intuition congruence. auto. auto. @@ -1808,8 +1808,8 @@ Proof. (* for value *) exploit tr_simple_lvalue; eauto. intros [SL1 [TY1 EV1]]. exploit tr_simple_lvalue. eauto. - eapply tr_expr_invariant with (le' := PTree.set t v1 le). eauto. - intros. apply PTree.gso. intuition congruence. + eapply tr_expr_invariant with (le' := ATree.set t v1 le). eauto. + intros. apply ATree.gso. intuition congruence. intros [SL2 [TY2 EV2]]. subst; simpl Kseqlist. econstructor; split. @@ -1818,16 +1818,16 @@ Proof. constructor. eapply step_make_assign; eauto. unfold transl_incrdecr. destruct id; simpl in H2. - econstructor. constructor. apply PTree.gss. constructor. + econstructor. constructor. apply ATree.gss. constructor. rewrite comp_env_preserved; simpl; eauto. - econstructor. constructor. apply PTree.gss. constructor. + econstructor. constructor. apply ATree.gss. constructor. rewrite comp_env_preserved; simpl; eauto. destruct id; auto. traceEq. econstructor. auto. apply S. apply tr_val_gen. auto. intros. econstructor; eauto. - rewrite H5; auto. apply PTree.gss. - intros. apply PTree.gso. intuition congruence. + rewrite H5; auto. apply ATree.gss. + intros. apply ATree.gso. intuition congruence. auto. auto. (* postincr stuck *) exploit tr_top_leftcontext; eauto. clear H11. @@ -1872,8 +1872,8 @@ Proof. econstructor. econstructor; eauto. rewrite <- TY1; eauto. traceEq. econstructor; eauto. change sl2 with (final For_val (Etempvar t (Csyntax.typeof r)) ++ sl2). apply S. - constructor. auto. intros. constructor. rewrite H2; auto. apply PTree.gss. - intros. apply PTree.gso. intuition congruence. + constructor. auto. intros. constructor. rewrite H2; auto. apply ATree.gss. + intros. apply ATree.gso. intuition congruence. auto. (* for effects *) econstructor; split. @@ -1890,8 +1890,8 @@ Proof. rewrite <- TY1; eauto. traceEq. econstructor; eauto. apply S. constructor; auto. - intros. constructor. rewrite H2. apply PTree.gss. auto. - intros. apply PTree.gso. congruence. + intros. constructor. rewrite H2. apply ATree.gss. auto. + intros. apply ATree.gso. congruence. auto. (* call *) @@ -1924,9 +1924,9 @@ Proof. constructor; auto. econstructor; eauto. intros. apply S. destruct dst'; constructor. - auto. intros. constructor. rewrite H5; auto. apply PTree.gss. - auto. intros. constructor. rewrite H5; auto. apply PTree.gss. - intros. apply PTree.gso. intuition congruence. + auto. intros. constructor. rewrite H5; auto. apply ATree.gss. + auto. intros. constructor. rewrite H5; auto. apply ATree.gss. + intros. apply ATree.gso. intuition congruence. auto. (* builtin *) @@ -1953,8 +1953,8 @@ Proof. traceEq. econstructor; eauto. change sl2 with (nil ++ sl2). apply S. - apply tr_val_gen. auto. intros. constructor. rewrite H2; auto. simpl. apply PTree.gss. - intros; simpl. apply PTree.gso. intuition congruence. + apply tr_val_gen. auto. intros. constructor. rewrite H2; auto. simpl. apply ATree.gss. + intros; simpl. apply ATree.gso. intuition congruence. auto. Qed. diff --git a/cfrontend/SimplExprspec.v b/cfrontend/SimplExprspec.v index 37e2cd9603..819b84be39 100644 --- a/cfrontend/SimplExprspec.v +++ b/cfrontend/SimplExprspec.v @@ -12,6 +12,7 @@ (** Relational specification of expression simplification. *) +Require Import String. Require Import Coqlib Maps Errors Integers Floats. Require Import AST Linking Memory. Require Import Ctypes Cop Csyntax Clight SimplExpr. @@ -65,14 +66,14 @@ Inductive tr_expr: temp_env -> destination -> Csyntax.expr -> list statement -> | tr_val_value: forall le v ty a tmp, typeof a = ty -> (forall tge e le' m, - (forall id, In id tmp -> le'!id = le!id) -> + (forall id, In id tmp -> le'$id = le$id) -> eval_expr tge e le' m a v) -> tr_expr le For_val (Csyntax.Eval v ty) nil a tmp | tr_val_set: forall le sd v ty a any tmp, typeof a = ty -> (forall tge e le' m, - (forall id, In id tmp -> le'!id = le!id) -> + (forall id, In id tmp -> le'$id = le$id) -> eval_expr tge e le' m a v) -> tr_expr le (For_set sd) (Csyntax.Eval v ty) (do_set sd a) any tmp @@ -327,16 +328,16 @@ Combined Scheme tr_expr_exprlist from tr_expr_ind2, tr_exprlist_ind2. Lemma tr_expr_invariant: forall le dst r sl a tmps, tr_expr le dst r sl a tmps -> - forall le', (forall x, In x tmps -> le'!x = le!x) -> + forall le', (forall x, In x tmps -> le'$x = le$x) -> tr_expr le' dst r sl a tmps with tr_exprlist_invariant: forall le rl sl al tmps, tr_exprlist le rl sl al tmps -> - forall le', (forall x, In x tmps -> le'!x = le!x) -> + forall le', (forall x, In x tmps -> le'$x = le$x) -> tr_exprlist le' rl sl al tmps. Proof. induction 1; intros; econstructor; eauto. - intros. apply H0. intros. transitivity (le'!id); auto. - intros. apply H0. auto. intros. transitivity (le'!id); auto. + intros. apply H0. intros. transitivity (le'$id); auto. + intros. apply H0. auto. intros. transitivity (le'$id); auto. induction 1; intros; econstructor; eauto. Qed. @@ -552,16 +553,70 @@ Ltac monadInv H := ((progress simpl in H) || unfold F in H); monadInv1 H end. +(** ** Injectivity of symbol generation *) + +Lemma app_one_inj s1 s2 c1 c2: + (s1 ++ String c1 "" = s2 ++ String c2 "")%string -> + s1 = s2 /\ c1 = c2. +Proof. + revert s2 c1 c2. + induction s1; intros. + - destruct s2; inv H; eauto. + induction s2; discriminate. + - destruct s2; inv H. + + induction s1; discriminate. + + apply IHs1 in H2. intuition congruence. +Qed. + +Lemma string_of_resid_inj: + forall x y, string_of_resid x = string_of_resid y -> x = y. +Proof. + intros x y H. + inv H. cut (N.pos x = N.pos y). { congruence. } + revert H1. generalize (N.pos x), (N.pos y). clear. + intros p. + functional induction (digits_of p); intros m Hm. + - apply Neqb_ok in e; subst. + rewrite digits_of_equation in Hm. + destruct (N.eqb_spec m 0); eauto. + destruct N.div_eucl. induction digits_of; discriminate. + - apply N.eqb_neq in e. + rewrite (digits_of_equation m) in Hm. + destruct (N.eqb_spec m 0); eauto. + destruct N.div_eucl. induction digits_of; discriminate. + destruct n as [|n]; try congruence. simpl in *. + destruct m as [|m]; try congruence. simpl in *. + pose proof (N.pos_div_eucl_remainder n 10) as Hrem. + pose proof (N.pos_div_eucl_remainder m 10) as Hrem'. + pose proof (N.pos_div_eucl_spec n 10) as Hspec. + pose proof (N.pos_div_eucl_spec m 10) as Hspec'. + rewrite e0 in Hspec, Hrem; subst. simpl in Hrem. + destruct (N.pos_div_eucl m 10) as [q' r']; subst. simpl in Hrem'. + cut (q = q' /\ r = r'). { xomega. } + clear - IHs Hrem Hrem' Hm. + apply app_one_inj in Hm. destruct Hm as [Hq Hr]. + split; eauto. + assert (10 <> 0)%N by congruence. specialize (Hrem H). specialize (Hrem' H). + cut (48 + r = 48 + r')%N. { xomega. } + rewrite <- (Ascii.N_ascii_embedding (48 + r)%N) by xomega. + rewrite <- (Ascii.N_ascii_embedding (48 + r')%N) by xomega. + congruence. +Qed. + (** ** Freshness and separation properties. *) Definition within (id: ident) (g1 g2: generator) : Prop := - Ple (gen_next g1) id /\ Plt id (gen_next g2). + exists p, + id = #(string_of_resid p) /\ + Ple (gen_next g1) p /\ Plt p (gen_next g2). Lemma gensym_within: forall ty g1 id g2 I, gensym ty g1 = Res id g2 I -> within id g1 g2. Proof. - intros. monadInv H. split. apply Ple_refl. apply Plt_succ. + intros. monadInv H. + exists (gen_next g1); intuition idtac. + apply Ple_refl. apply Plt_succ. Qed. Lemma within_widen: @@ -571,7 +626,8 @@ Lemma within_widen: Ple (gen_next g2) (gen_next g2') -> within id g1' g2'. Proof. - intros. destruct H. split. + intros. destruct H as (? & ? & ?). + exists x; intuition idtac. eapply Ple_trans; eauto. eapply Plt_Ple_trans; eauto. Qed. @@ -614,21 +670,26 @@ Lemma contained_disjoint: contained l1 g1 g2 -> contained l2 g2 g3 -> list_disjoint l1 l2. Proof. intros; red; intros. red; intro; subst y. - exploit H; eauto. intros [A B]. exploit H0; eauto. intros [Csyntax D]. - elim (Plt_strict x). apply Plt_Ple_trans with (gen_next g2); auto. + exploit H; eauto. intros (p & Hp & A & B). + exploit H0; eauto. intros (q & Hq & Csyntax & D). + assert (p = q) by (apply string_of_resid_inj, ident_of_string_inj; congruence). + subst q. elim (Plt_strict p). apply Plt_Ple_trans with (gen_next g2); auto. Qed. Lemma contained_notin: forall g1 l g2 id g3, contained l g1 g2 -> within id g2 g3 -> ~In id l. Proof. - intros; red; intros. exploit H; eauto. intros [Csyntax D]. destruct H0 as [A B]. - elim (Plt_strict id). apply Plt_Ple_trans with (gen_next g2); auto. + intros; red; intros. + exploit H; eauto. intros (p & Hp & Csyntax & D). + destruct H0 as (q & Hq & A & B). + assert (p = q) by (apply string_of_resid_inj, ident_of_string_inj; congruence). + subst q. elim (Plt_strict p). apply Plt_Ple_trans with (gen_next g2); auto. Qed. Definition dest_below (dst: destination) (g: generator) : Prop := match dst with - | For_set sd => Plt (sd_temp sd) g.(gen_next) + | For_set sd => exists p, sd_temp sd = #(string_of_resid p) /\ Plt p g.(gen_next) | _ => True end. @@ -642,7 +703,7 @@ Lemma dest_for_set_seqbool_val: forall tmp ty g1 g2, within tmp g1 g2 -> dest_below (For_set (sd_seqbool_val tmp ty)) g2. Proof. - intros. destruct H. simpl. auto. + intros. destruct H as (p & Hp & A & B). simpl. eauto. Qed. Lemma dest_for_set_seqbool_set: @@ -654,27 +715,30 @@ Qed. Lemma dest_for_set_condition_val: forall tmp tycast ty g1 g2, within tmp g1 g2 -> dest_below (For_set (SDbase tycast ty tmp)) g2. Proof. - intros. destruct H. simpl. auto. + intros. destruct H as (p & Hp & A & B). simpl. eauto. Qed. Lemma dest_for_set_condition_set: forall sd tmp tycast ty g1 g2, dest_below (For_set sd) g2 -> within tmp g1 g2 -> dest_below (For_set (SDcons tycast ty tmp sd)) g2. Proof. - intros. destruct H0. simpl. auto. + intros. destruct H0 as (p & Hp & A & B). simpl. eauto. Qed. Lemma sd_temp_notin: forall sd g1 g2 l, dest_below (For_set sd) g1 -> contained l g1 g2 -> ~In (sd_temp sd) l. Proof. - intros. simpl in H. red; intros. exploit H0; eauto. intros [A B]. - elim (Plt_strict (sd_temp sd)). apply Plt_Ple_trans with (gen_next g1); auto. + intros. destruct H as (p & Hp & H). red; intros. + exploit H0; eauto. intros (q & Hq & A & B). + assert (p = q) by (apply string_of_resid_inj, ident_of_string_inj; congruence). + subst q. elim (Plt_strict p). apply Plt_Ple_trans with (gen_next g1); auto. Qed. Lemma dest_below_le: forall dst g1 g2, dest_below dst g1 -> Ple g1.(gen_next) g2.(gen_next) -> dest_below dst g2. Proof. - intros. destruct dst; simpl in *; auto. eapply Plt_Ple_trans; eauto. + intros. destruct dst; simpl in *; auto. + destruct H as (p & Hp & H). eauto using Plt_Ple_trans. Qed. Hint Resolve gensym_within within_widen contained_widen diff --git a/cfrontend/SimplLocals.v b/cfrontend/SimplLocals.v index b142d3cc51..e7a47c08da 100644 --- a/cfrontend/SimplLocals.v +++ b/cfrontend/SimplLocals.v @@ -23,7 +23,7 @@ Require Compopts. Open Scope error_monad_scope. Open Scope string_scope. -Module VSet := FSetAVL.Make(OrderedPositive). +Module VSet := FSetAVL.Make(OrderedIdent). (** The set of local variables that can be lifted to temporaries, because they are scalar and their address is not taken. *) diff --git a/cfrontend/SimplLocalsproof.v b/cfrontend/SimplLocalsproof.v index 7af499f487..b627267c16 100644 --- a/cfrontend/SimplLocalsproof.v +++ b/cfrontend/SimplLocalsproof.v @@ -78,24 +78,24 @@ Qed. Inductive match_var (f: meminj) (cenv: compilenv) (e: env) (m: mem) (te: env) (tle: temp_env) (id: ident) : Prop := | match_var_lifted: forall b ty chunk v tv - (ENV: e!id = Some(b, ty)) - (TENV: te!id = None) + (ENV: e$id = Some(b, ty)) + (TENV: te$id = None) (LIFTED: VSet.mem id cenv = true) (MAPPED: f b = None) (MODE: access_mode ty = By_value chunk) (LOAD: Mem.load chunk m b 0 = Some v) - (TLENV: tle!(id) = Some tv) + (TLENV: tle$(id) = Some tv) (VINJ: Val.inject f v tv), match_var f cenv e m te tle id | match_var_not_lifted: forall b ty b' - (ENV: e!id = Some(b, ty)) - (TENV: te!id = Some(b', ty)) + (ENV: e$id = Some(b, ty)) + (TENV: te$id = Some(b', ty)) (LIFTED: VSet.mem id cenv = false) (MAPPED: f b = Some(b', 0)), match_var f cenv e m te tle id | match_var_not_local: forall - (ENV: e!id = None) - (TENV: te!id = None) + (ENV: e$id = None) + (TENV: te$id = None) (LIFTED: VSet.mem id cenv = false), match_var f cenv e m te tle id. @@ -107,21 +107,21 @@ Record match_envs (f: meminj) (cenv: compilenv) forall id, match_var f cenv e m te tle id; me_temps: forall id v, - le!id = Some v -> - (exists tv, tle!id = Some tv /\ Val.inject f v tv) + le$id = Some v -> + (exists tv, tle$id = Some tv /\ Val.inject f v tv) /\ (VSet.mem id cenv = true -> v = Vundef); me_inj: - forall id1 b1 ty1 id2 b2 ty2, e!id1 = Some(b1, ty1) -> e!id2 = Some(b2, ty2) -> id1 <> id2 -> b1 <> b2; + forall id1 b1 ty1 id2 b2 ty2, e$id1 = Some(b1, ty1) -> e$id2 = Some(b2, ty2) -> id1 <> id2 -> b1 <> b2; me_range: - forall id b ty, e!id = Some(b, ty) -> Ple lo b /\ Plt b hi; + forall id b ty, e$id = Some(b, ty) -> Ple lo b /\ Plt b hi; me_trange: - forall id b ty, te!id = Some(b, ty) -> Ple tlo b /\ Plt b thi; + forall id b ty, te$id = Some(b, ty) -> Ple tlo b /\ Plt b thi; me_mapped: forall id b' ty, - te!id = Some(b', ty) -> exists b, f b = Some(b', 0) /\ e!id = Some(b, ty); + te$id = Some(b', ty) -> exists b, f b = Some(b', 0) /\ e$id = Some(b, ty); me_flat: forall id b' ty b delta, - te!id = Some(b', ty) -> f b = Some(b', delta) -> e!id = Some(b, ty) /\ delta = 0; + te$id = Some(b', ty) -> f b = Some(b', delta) -> e$id = Some(b, ty) /\ delta = 0; me_incr: Ple lo hi; me_tincr: @@ -288,7 +288,7 @@ Qed. Lemma step_Sdebug_temp: forall f id ty k e le m v, - le!id = Some v -> + le$id = Some v -> val_casted v ty -> step2 tge (State f (Sdebug_temp id ty) k e le m) E0 (State f Sskip k e le m). @@ -300,7 +300,7 @@ Qed. Lemma step_Sdebug_var: forall f id ty k e le m b, - e!id = Some(b, ty) -> + e$id = Some(b, ty) -> step2 tge (State f (Sdebug_var id ty) k e le m) E0 (State f Sskip k e le m). Proof. @@ -315,25 +315,25 @@ Lemma step_Sset_debug: eval_expr tge e le m a v -> sem_cast v (typeof a) ty m = Some v' -> plus step2 tge (State f (Sset_debug id ty a) k e le m) - E0 (State f Sskip k e (PTree.set id v' le) m). + E0 (State f Sskip k e (ATree.set id v' le) m). Proof. intros; unfold Sset_debug. assert (forall k, step2 tge (State f (Sset id (make_cast a ty)) k e le m) - E0 (State f Sskip k e (PTree.set id v' le) m)). + E0 (State f Sskip k e (ATree.set id v' le) m)). { intros. apply step_set. eapply make_cast_correct; eauto. } destruct (Compopts.debug tt). - eapply plus_left. constructor. eapply star_left. apply H1. eapply star_left. constructor. apply star_one. apply step_Sdebug_temp with (v := v'). - apply PTree.gss. eapply cast_val_is_casted; eauto. + apply ATree.gss. eapply cast_val_is_casted; eauto. reflexivity. reflexivity. reflexivity. - apply plus_one. apply H1. Qed. Lemma step_add_debug_vars: forall f s e le m vars k, - (forall id ty, In (id, ty) vars -> exists b, e!id = Some (b, ty)) -> + (forall id ty, In (id, ty) vars -> exists b, e$id = Some (b, ty)) -> star step2 tge (State f (add_debug_vars vars s) k e le m) E0 (State f s k e le m). Proof. @@ -354,13 +354,13 @@ Remark bind_parameter_temps_inv: forall id params args le le', bind_parameter_temps params args le = Some le' -> ~In id (var_names params) -> - le'!id = le!id. + le'$id = le$id. Proof. induction params; simpl; intros. destruct args; inv H. auto. destruct a as [id1 ty1]. destruct args; try discriminate. - transitivity ((PTree.set id1 v le)!id). - eapply IHparams; eauto. apply PTree.gso. intuition. + transitivity ((ATree.set id1 v le)$id). + eapply IHparams; eauto. apply ATree.gso. intuition. Qed. Lemma step_add_debug_params: @@ -374,7 +374,7 @@ Proof. unfold add_debug_params. destruct (Compopts.debug tt). - induction params as [ | [id ty] params ]; simpl; intros until le1; intros NR CAST BIND; inv CAST; inv NR. + apply star_refl. - + assert (le!id = Some a1). { erewrite bind_parameter_temps_inv by eauto. apply PTree.gss. } + + assert (le$id = Some a1). { erewrite bind_parameter_temps_inv by eauto. apply ATree.gss. } eapply star_left. constructor. eapply star_left. eapply step_Sdebug_temp; eauto. eapply star_left. constructor. @@ -388,31 +388,31 @@ Qed. Lemma match_envs_assign_lifted: forall f cenv e le m lo hi te tle tlo thi b ty v m' id tv, match_envs f cenv e le m lo hi te tle tlo thi -> - e!id = Some(b, ty) -> + e$id = Some(b, ty) -> val_casted v ty -> Val.inject f v tv -> assign_loc ge ty m b Ptrofs.zero v m' -> VSet.mem id cenv = true -> - match_envs f cenv e le m' lo hi te (PTree.set id tv tle) tlo thi. + match_envs f cenv e le m' lo hi te (ATree.set id tv tle) tlo thi. Proof. intros. destruct H. generalize (me_vars0 id); intros MV; inv MV; try congruence. rewrite ENV in H0; inv H0. inv H3; try congruence. unfold Mem.storev in H0. rewrite Ptrofs.unsigned_zero in H0. constructor; eauto; intros. (* vars *) - destruct (peq id0 id). subst id0. + destruct (ATree.elt_eq id0 id). subst id0. eapply match_var_lifted with (v := v); eauto. exploit Mem.load_store_same; eauto. erewrite val_casted_load_result; eauto. - apply PTree.gss. + apply ATree.gss. generalize (me_vars0 id0); intros MV; inv MV. eapply match_var_lifted; eauto. rewrite <- LOAD0. eapply Mem.load_store_other; eauto. - rewrite PTree.gso; auto. + rewrite ATree.gso; auto. eapply match_var_not_lifted; eauto. eapply match_var_not_local; eauto. (* temps *) exploit me_temps0; eauto. intros [[tv1 [A B]] C]. split; auto. - rewrite PTree.gsspec. destruct (peq id0 id). + rewrite ATree.gsspec. destruct (ATree.elt_eq id0 id). subst id0. exists tv; split; auto. rewrite C; auto. exists tv1; auto. Qed. @@ -424,18 +424,18 @@ Lemma match_envs_set_temp: match_envs f cenv e le m lo hi te tle tlo thi -> Val.inject f v tv -> check_temp cenv id = OK x -> - match_envs f cenv e (PTree.set id v le) m lo hi te (PTree.set id tv tle) tlo thi. + match_envs f cenv e (ATree.set id v le) m lo hi te (ATree.set id tv tle) tlo thi. Proof. intros. unfold check_temp in H1. destruct (VSet.mem id cenv) eqn:?; monadInv H1. destruct H. constructor; eauto; intros. (* vars *) generalize (me_vars0 id0); intros MV; inv MV. - eapply match_var_lifted; eauto. rewrite PTree.gso. eauto. congruence. + eapply match_var_lifted; eauto. rewrite ATree.gso. eauto. congruence. eapply match_var_not_lifted; eauto. eapply match_var_not_local; eauto. (* temps *) - rewrite PTree.gsspec in *. destruct (peq id0 id). + rewrite ATree.gsspec in *. destruct (ATree.elt_eq id0 id). inv H. split. exists tv; auto. intros; congruence. eapply me_temps0; eauto. Qed. @@ -457,7 +457,7 @@ Qed. Lemma match_envs_temps_exten: forall f cenv e le m lo hi te tle tlo thi tle', match_envs f cenv e le m lo hi te tle tlo thi -> - (forall id, tle'!id = tle!id) -> + (forall id, tle'$id = tle$id) -> match_envs f cenv e le m lo hi te tle' tlo thi. Proof. intros. destruct H. constructor; auto; intros. @@ -475,17 +475,17 @@ Qed. Lemma match_envs_change_temp: forall f cenv e le m lo hi te tle tlo thi id v, match_envs f cenv e le m lo hi te tle tlo thi -> - le!id = None -> VSet.mem id cenv = false -> - match_envs f cenv e le m lo hi te (PTree.set id v tle) tlo thi. + le$id = None -> VSet.mem id cenv = false -> + match_envs f cenv e le m lo hi te (ATree.set id v tle) tlo thi. Proof. intros. destruct H. constructor; auto; intros. (* vars *) generalize (me_vars0 id0); intros MV; inv MV. - eapply match_var_lifted; eauto. rewrite PTree.gso; auto. congruence. + eapply match_var_lifted; eauto. rewrite ATree.gso; auto. congruence. eapply match_var_not_lifted; eauto. eapply match_var_not_local; eauto. (* temps *) - rewrite PTree.gso. eauto. congruence. + rewrite ATree.gso. eauto. congruence. Qed. (** Properties of [cenv_for]. *) @@ -596,12 +596,12 @@ Qed. Lemma alloc_variables_range: forall ge id b ty e m vars e' m', alloc_variables ge e m vars e' m' -> - e'!id = Some(b, ty) -> e!id = Some(b, ty) \/ Ple (Mem.nextblock m) b /\ Plt b (Mem.nextblock m'). + e'$id = Some(b, ty) -> e$id = Some(b, ty) \/ Ple (Mem.nextblock m) b /\ Plt b (Mem.nextblock m'). Proof. induction 1; intros. auto. - exploit IHalloc_variables; eauto. rewrite PTree.gsspec. intros [A|A]. - destruct (peq id id0). inv A. + exploit IHalloc_variables; eauto. rewrite ATree.gsspec. intros [A|A]. + destruct (ATree.elt_eq id id0). inv A. right. exploit Mem.alloc_result; eauto. exploit Mem.nextblock_alloc; eauto. generalize (alloc_variables_nextblock _ _ _ _ _ _ H0). intros A B C. subst b. split. apply Ple_refl. eapply Pos.lt_le_trans; eauto. rewrite B. apply Plt_succ. @@ -612,20 +612,20 @@ Qed. Lemma alloc_variables_injective: forall ge id1 b1 ty1 id2 b2 ty2 e m vars e' m', alloc_variables ge e m vars e' m' -> - (e!id1 = Some(b1, ty1) -> e!id2 = Some(b2, ty2) -> id1 <> id2 -> b1 <> b2) -> - (forall id b ty, e!id = Some(b, ty) -> Plt b (Mem.nextblock m)) -> - (e'!id1 = Some(b1, ty1) -> e'!id2 = Some(b2, ty2) -> id1 <> id2 -> b1 <> b2). + (e$id1 = Some(b1, ty1) -> e$id2 = Some(b2, ty2) -> id1 <> id2 -> b1 <> b2) -> + (forall id b ty, e$id = Some(b, ty) -> Plt b (Mem.nextblock m)) -> + (e'$id1 = Some(b1, ty1) -> e'$id2 = Some(b2, ty2) -> id1 <> id2 -> b1 <> b2). Proof. induction 1; intros. eauto. eapply IHalloc_variables; eauto. - repeat rewrite PTree.gsspec; intros. - destruct (peq id1 id); destruct (peq id2 id). + repeat rewrite ATree.gsspec; intros. + destruct (ATree.elt_eq id1 id); destruct (ATree.elt_eq id2 id). congruence. inv H6. exploit Mem.alloc_result; eauto. exploit H2; eauto. unfold block; xomega. inv H7. exploit Mem.alloc_result; eauto. exploit H2; eauto. unfold block; xomega. eauto. - intros. rewrite PTree.gsspec in H6. destruct (peq id0 id). inv H6. + intros. rewrite ATree.gsspec in H6. destruct (ATree.elt_eq id0 id). inv H6. exploit Mem.alloc_result; eauto. exploit Mem.nextblock_alloc; eauto. unfold block; xomega. exploit H2; eauto. exploit Mem.nextblock_alloc; eauto. unfold block; xomega. Qed. @@ -643,14 +643,14 @@ Lemma match_alloc_variables: /\ (forall b, Mem.valid_block m b -> j' b = j b) /\ (forall b b' delta, j' b = Some(b', delta) -> Mem.valid_block tm b' -> j' b = j b) /\ (forall b b' delta, j' b = Some(b', delta) -> ~Mem.valid_block tm b' -> - exists id, exists ty, e'!id = Some(b, ty) /\ te'!id = Some(b', ty) /\ delta = 0) + exists id, exists ty, e'$id = Some(b, ty) /\ te'$id = Some(b', ty) /\ delta = 0) /\ (forall id ty, In (id, ty) vars -> exists b, - e'!id = Some(b, ty) + e'$id = Some(b, ty) /\ if VSet.mem id cenv - then te'!id = te!id /\ j' b = None - else exists tb, te'!id = Some(tb, ty) /\ j' b = Some(tb, 0)) - /\ (forall id, ~In id (var_names vars) -> e'!id = e!id /\ te'!id = te!id). + then te'$id = te$id /\ j' b = None + else exists tb, te'$id = Some(tb, ty) /\ j' b = Some(tb, 0)) + /\ (forall id, ~In id (var_names vars) -> e'$id = e$id /\ te'$id = te$id). Proof. induction 1; intros. (* base case *) @@ -686,19 +686,19 @@ Proof. destruct H1. congruence. elim H5. unfold var_names. change id with (fst (id, ty0)). apply in_map; auto. subst ty0. exploit P; eauto. intros [X Y]. rewrite Heqb. rewrite X. rewrite Y. - exists b1. split. apply PTree.gss. + exists b1. split. apply ATree.gss. split. auto. rewrite M. auto. eapply Mem.valid_new_block; eauto. (* other vars *) eapply O; eauto. destruct H1. congruence. auto. intros. exploit (P id0). tauto. intros [X Y]. rewrite X; rewrite Y. - split; auto. apply PTree.gso. intuition. + split; auto. apply ATree.gso. intuition. (* variable is not lifted out of memory *) exploit Mem.alloc_parallel_inject. eauto. eauto. apply Z.le_refl. apply Z.le_refl. intros [j1 [tm1 [tb1 [A [B [C [D E]]]]]]]. - exploit IHalloc_variables; eauto. instantiate (1 := PTree.set id (tb1, ty) te). + exploit IHalloc_variables; eauto. instantiate (1 := ATree.set id (tb1, ty) te). intros [j' [te' [tm' [J [K [L [M [N [Q [O P]]]]]]]]]]. exists j'; exists te'; exists tm'. split. simpl. econstructor; eauto. rewrite comp_env_preserved; auto. @@ -714,7 +714,7 @@ Proof. subst b'. rewrite (N _ _ _ H1) in H1. destruct (eq_block b b1). subst b. rewrite D in H1; inv H1. exploit (P id); auto. intros [X Y]. exists id; exists ty. - rewrite X; rewrite Y. repeat rewrite PTree.gss. auto. + rewrite X; rewrite Y. repeat rewrite ATree.gss. auto. rewrite E in H1; auto. elim H3. eapply Mem.mi_mappedblocks; eauto. eapply Mem.valid_new_block; eauto. eapply Q; eauto. unfold Mem.valid_block in *. @@ -727,15 +727,15 @@ Proof. destruct H1. congruence. elim H5. unfold var_names. change id with (fst (id, ty0)). apply in_map; auto. subst ty0. exploit P; eauto. intros [X Y]. rewrite Heqb. rewrite X. rewrite Y. - exists b1. split. apply PTree.gss. + exists b1. split. apply ATree.gss. exists tb1; split. - apply PTree.gss. + apply ATree.gss. rewrite M. auto. eapply Mem.valid_new_block; eauto. (* other vars *) exploit (O id0 ty0). destruct H1. congruence. auto. - rewrite PTree.gso; auto. + rewrite ATree.gso; auto. intros. exploit (P id0). tauto. intros [X Y]. rewrite X; rewrite Y. - split; apply PTree.gso; intuition. + split; apply ATree.gso; intuition. Qed. Lemma alloc_variables_load: @@ -765,7 +765,7 @@ Qed. Definition env_initial_value (e: env) (m: mem) := forall id b ty chunk, - e!id = Some(b, ty) -> access_mode ty = By_value chunk -> Mem.load chunk m b 0 = Some Vundef. + e$id = Some(b, ty) -> access_mode ty = By_value chunk -> Mem.load chunk m b 0 = Some Vundef. Lemma alloc_variables_initial_value: forall e m vars e' m', @@ -775,8 +775,8 @@ Lemma alloc_variables_initial_value: Proof. induction 1; intros. auto. - apply IHalloc_variables. red; intros. rewrite PTree.gsspec in H2. - destruct (peq id0 id). inv H2. + apply IHalloc_variables. red; intros. rewrite ATree.gsspec in H2. + destruct (ATree.elt_eq id0 id). inv H2. eapply Mem.load_alloc_same'; eauto. omega. rewrite Z.add_0_l. eapply sizeof_by_value; eauto. apply Z.divide_0_r. @@ -784,20 +784,20 @@ Proof. Qed. Lemma create_undef_temps_charact: - forall id ty vars, In (id, ty) vars -> (create_undef_temps vars)!id = Some Vundef. + forall id ty vars, In (id, ty) vars -> (create_undef_temps vars)$id = Some Vundef. Proof. induction vars; simpl; intros. contradiction. - destruct H. subst a. apply PTree.gss. - destruct a as [id1 ty1]. rewrite PTree.gsspec. destruct (peq id id1); auto. + destruct H. subst a. apply ATree.gss. + destruct a as [id1 ty1]. rewrite ATree.gsspec. destruct (ATree.elt_eq id id1); auto. Qed. Lemma create_undef_temps_inv: - forall vars id v, (create_undef_temps vars)!id = Some v -> v = Vundef /\ In id (var_names vars). + forall vars id v, (create_undef_temps vars)$id = Some v -> v = Vundef /\ In id (var_names vars). Proof. induction vars; simpl; intros. - rewrite PTree.gempty in H; congruence. - destruct a as [id1 ty1]. rewrite PTree.gsspec in H. destruct (peq id id1). + rewrite ATree.gempty in H; congruence. + destruct a as [id1 ty1]. rewrite ATree.gsspec in H. destruct (ATree.elt_eq id id1). inv H. auto. exploit IHvars; eauto. tauto. Qed. @@ -805,12 +805,12 @@ Qed. Lemma create_undef_temps_exten: forall id l1 l2, (In id (var_names l1) <-> In id (var_names l2)) -> - (create_undef_temps l1)!id = (create_undef_temps l2)!id. + (create_undef_temps l1)$id = (create_undef_temps l2)$id. Proof. assert (forall id l1 l2, (In id (var_names l1) -> In id (var_names l2)) -> - (create_undef_temps l1)!id = None \/ (create_undef_temps l1)!id = (create_undef_temps l2)!id). - intros. destruct ((create_undef_temps l1)!id) as [v1|] eqn:?; auto. + (create_undef_temps l1)$id = None \/ (create_undef_temps l1)$id = (create_undef_temps l2)$id). + intros. destruct ((create_undef_temps l1)$id) as [v1|] eqn:?; auto. exploit create_undef_temps_inv; eauto. intros [A B]. subst v1. exploit list_in_map_inv. unfold var_names in H. apply H. eexact B. intros [[id1 ty1] [P Q]]. simpl in P; subst id1. @@ -867,8 +867,8 @@ Qed. Lemma create_undef_temps_lifted: forall id f, ~ In id (var_names (fn_params f)) -> - (create_undef_temps (add_lifted (cenv_for f) (fn_vars f) (fn_temps f))) ! id = - (create_undef_temps (add_lifted (cenv_for f) (fn_params f ++ fn_vars f) (fn_temps f))) ! id. + (create_undef_temps (add_lifted (cenv_for f) (fn_vars f) (fn_temps f))) $ id = + (create_undef_temps (add_lifted (cenv_for f) (fn_params f ++ fn_vars f) (fn_temps f))) $ id. Proof. intros. apply create_undef_temps_exten. unfold add_lifted. rewrite filter_app. @@ -918,7 +918,7 @@ Theorem match_envs_alloc_variables: /\ inject_incr j j' /\ (forall b, Mem.valid_block m b -> j' b = j b) /\ (forall b b' delta, j' b = Some(b', delta) -> Mem.valid_block tm b' -> j' b = j b) - /\ (forall id ty, In (id, ty) vars -> VSet.mem id cenv = false -> exists b, te!id = Some(b, ty)). + /\ (forall id ty, In (id, ty) vars -> VSet.mem id cenv = false -> exists b, te$id = Some(b, ty)). Proof. intros. exploit (match_alloc_variables cenv); eauto. instantiate (1 := empty_env). @@ -935,9 +935,9 @@ Proof. (* local var, lifted *) destruct R as [U V]. exploit H2; eauto. intros [chunk X]. eapply match_var_lifted with (v := Vundef) (tv := Vundef); eauto. - rewrite U; apply PTree.gempty. + rewrite U; apply ATree.gempty. eapply alloc_variables_initial_value; eauto. - red. unfold empty_env; intros. rewrite PTree.gempty in H4; congruence. + red. unfold empty_env; intros. rewrite ATree.gempty in H4; congruence. apply create_undef_temps_charact with ty. unfold add_lifted. apply in_or_app. left. rewrite filter_In. auto. @@ -945,7 +945,7 @@ Proof. destruct R as [tb [U V]]. eapply match_var_not_lifted; eauto. (* non-local var *) - exploit G; eauto. unfold empty_env. rewrite PTree.gempty. intros [U V]. + exploit G; eauto. unfold empty_env. rewrite ATree.gempty. intros [U V]. eapply match_var_not_local; eauto. destruct (VSet.mem id cenv) eqn:?; auto. elim n; eauto. @@ -960,17 +960,17 @@ Proof. (* injective *) eapply alloc_variables_injective. eexact H. - rewrite PTree.gempty. congruence. - intros. rewrite PTree.gempty in H7. congruence. + rewrite ATree.gempty. congruence. + intros. rewrite ATree.gempty in H7. congruence. eauto. eauto. auto. (* range *) exploit alloc_variables_range. eexact H. eauto. - rewrite PTree.gempty. intuition congruence. + rewrite ATree.gempty. intuition congruence. (* trange *) exploit alloc_variables_range. eexact A. eauto. - rewrite PTree.gempty. intuition congruence. + rewrite ATree.gempty. intuition congruence. (* mapped *) destruct (In_dec ident_eq id (var_names vars)). @@ -978,20 +978,20 @@ Proof. intros [[id' ty'] [EQ IN]]; simpl in EQ; subst id'. exploit F; eauto. intros [b [P Q]]. destruct (VSet.mem id cenv). - rewrite PTree.gempty in Q. destruct Q; congruence. + rewrite ATree.gempty in Q. destruct Q; congruence. destruct Q as [tb [U V]]. exists b; split; congruence. - exploit G; eauto. rewrite PTree.gempty. intuition congruence. + exploit G; eauto. rewrite ATree.gempty. intuition congruence. (* flat *) exploit alloc_variables_range. eexact A. eauto. - rewrite PTree.gempty. intros [P|P]. congruence. + rewrite ATree.gempty. intros [P|P]. congruence. exploit K; eauto. unfold Mem.valid_block. xomega. intros [id0 [ty0 [U [V W]]]]. split; auto. destruct (ident_eq id id0). congruence. assert (b' <> b'). eapply alloc_variables_injective with (e' := te) (id1 := id) (id2 := id0); eauto. - rewrite PTree.gempty; congruence. - intros until ty1; rewrite PTree.gempty; congruence. + rewrite ATree.gempty; congruence. + intros until ty1; rewrite ATree.gempty; congruence. congruence. (* incr *) @@ -1098,8 +1098,8 @@ Theorem store_params_correct: Val.inject_list j args targs -> match_envs j cenv e le m lo hi te tle1 tlo thi -> Mem.inject j m tm -> - (forall id, ~In id (var_names params) -> tle2!id = tle1!id) -> - (forall id, In id (var_names params) -> le!id = None) -> + (forall id, ~In id (var_names params) -> tle2$id = tle1$id) -> + (forall id, In id (var_names params) -> le$id = None) -> exists tle, exists tm', star step2 tge (State f (store_params cenv params s) k te tle tm) E0 (State f s k te tle tm') @@ -1117,25 +1117,25 @@ Proof. exploit me_vars; eauto. instantiate (1 := id); intros MV. destruct (VSet.mem id cenv) eqn:?. (* lifted to temp *) - eapply IHbind_parameters with (tle1 := PTree.set id v' tle1); eauto. + eapply IHbind_parameters with (tle1 := ATree.set id v' tle1); eauto. eapply match_envs_assign_lifted; eauto. inv MV; try congruence. rewrite ENV in H; inv H. inv H0; try congruence. unfold Mem.storev in H2. eapply Mem.store_unmapped_inject; eauto. - intros. repeat rewrite PTree.gsspec. destruct (peq id0 id). auto. + intros. repeat rewrite ATree.gsspec. destruct (ATree.elt_eq id0 id). auto. apply TLE. intuition. (* still in memory *) inv MV; try congruence. rewrite ENV in H; inv H. exploit assign_loc_inject; eauto. intros [tm1 [A [B C]]]. exploit IHbind_parameters. eauto. eauto. eauto. - instantiate (1 := PTree.set id v' tle1). + instantiate (1 := ATree.set id v' tle1). apply match_envs_change_temp. eapply match_envs_invariant; eauto. apply LE; auto. auto. eauto. - instantiate (1 := PTree.set id v' tle2). - intros. repeat rewrite PTree.gsspec. destruct (peq id0 id). auto. + instantiate (1 := ATree.set id v' tle2). + intros. repeat rewrite ATree.gsspec. destruct (ATree.elt_eq id0 id). auto. apply TLE. intuition. intros. apply LE. auto. instantiate (1 := s). @@ -1146,7 +1146,7 @@ Proof. eapply star_left. econstructor. eapply eval_Evar_local. eauto. eapply eval_Etempvar. erewrite bind_parameter_temps_inv; eauto. - apply PTree.gss. + apply ATree.gss. simpl. instantiate (1 := v'). apply cast_val_casted. eapply val_casted_inject with (v := v1); eauto. simpl. eexact A. @@ -1168,7 +1168,7 @@ Qed. Lemma bind_parameters_load: forall ge e chunk b ofs, - (forall id b' ty, e!id = Some(b', ty) -> b <> b') -> + (forall id b' ty, e$id = Some(b', ty) -> b <> b') -> forall m params args m', bind_parameters ge e m params args m' -> Mem.load chunk m' b ofs = Mem.load chunk m b ofs. @@ -1187,7 +1187,7 @@ Qed. Lemma free_blocks_of_env_perm_1: forall ce m e m' id b ty ofs k p, Mem.free_list m (blocks_of_env ce e) = Some m' -> - e!id = Some(b, ty) -> + e$id = Some(b, ty) -> Mem.perm m' b ofs k p -> 0 <= ofs < sizeof ce ty -> False. @@ -1195,7 +1195,7 @@ Proof. intros. exploit Mem.perm_free_list; eauto. intros [A B]. apply B with 0 (sizeof ce ty); auto. unfold blocks_of_env. change (b, 0, sizeof ce ty) with (block_of_binding ce (id, (b, ty))). - apply in_map. apply PTree.elements_correct. auto. + apply in_map. apply ATree.elements_correct. auto. Qed. Lemma free_list_perm': @@ -1215,12 +1215,12 @@ Qed. Lemma free_blocks_of_env_perm_2: forall ce m e m' id b ty, Mem.free_list m (blocks_of_env ce e) = Some m' -> - e!id = Some(b, ty) -> + e$id = Some(b, ty) -> Mem.range_perm m b 0 (sizeof ce ty) Cur Freeable. Proof. intros. eapply free_list_perm'; eauto. unfold blocks_of_env. change (b, 0, sizeof ce ty) with (block_of_binding ce (id, (b, ty))). - apply in_map. apply PTree.elements_correct. auto. + apply in_map. apply ATree.elements_correct. auto. Qed. Fixpoint freelist_no_overlap (l: list (block * Z * Z)) : Prop := @@ -1253,10 +1253,10 @@ Lemma blocks_of_env_no_overlap: match_envs j cenv e le m lo hi te tle tlo thi -> Mem.inject j m tm -> (forall id b ty, - e!id = Some(b, ty) -> Mem.range_perm m b 0 (sizeof ge ty) Cur Freeable) -> + e$id = Some(b, ty) -> Mem.range_perm m b 0 (sizeof ge ty) Cur Freeable) -> forall l, list_norepet (List.map fst l) -> - (forall id bty, In (id, bty) l -> te!id = Some bty) -> + (forall id bty, In (id, bty) l -> te$id = Some bty) -> freelist_no_overlap (List.map (block_of_binding ge) l). Proof. intros until tm; intros ME MINJ PERMS. induction l; simpl; intros. @@ -1265,8 +1265,8 @@ Proof. + apply IHl; auto. + intros. exploit list_in_map_inv; eauto. intros [[id' [b'' ty']] [A B]]. simpl in A. inv A. rename b'' into b'. - assert (TE: te!id = Some(b, ty)) by eauto. - assert (TE': te!id' = Some(b', ty')) by eauto. + assert (TE: te$id = Some(b, ty)) by eauto. + assert (TE': te$id' = Some(b', ty')) by eauto. exploit me_mapped. eauto. eexact TE. intros [b0 [INJ E]]. exploit me_mapped. eauto. eexact TE'. intros [b0' [INJ' E']]. destruct (zle (sizeof ge0 ty) 0); auto. @@ -1324,7 +1324,7 @@ Local Opaque ge tge. intros. unfold blocks_of_env in H2. exploit list_in_map_inv; eauto. intros [[id [b' ty]] [EQ IN]]. unfold block_of_binding in EQ; inv EQ. - exploit me_mapped; eauto. eapply PTree.elements_complete; eauto. + exploit me_mapped; eauto. eapply ATree.elements_complete; eauto. intros [b [A B]]. change 0 with (0 + 0). replace (sizeof ge ty) with (sizeof ge ty + 0) by omega. eapply Mem.range_perm_inject; eauto. @@ -1332,8 +1332,8 @@ Local Opaque ge tge. - (* no overlap *) unfold blocks_of_env; eapply blocks_of_env_no_overlap; eauto. intros. eapply free_blocks_of_env_perm_2; eauto. - apply PTree.elements_keys_norepet. - intros. apply PTree.elements_complete; auto. + apply ATree.elements_keys_norepet. + intros. apply ATree.elements_complete; auto. } destruct X as [tm' FREE]. exists tm'; split; auto. @@ -1341,7 +1341,7 @@ Local Opaque ge tge. eapply Mem.free_list_left_inject; eauto. intros. unfold blocks_of_env in H3. exploit list_in_map_inv; eauto. intros [[id [b' ty]] [EQ IN]]. unfold block_of_binding in EQ. inv EQ. - exploit me_flat; eauto. apply PTree.elements_complete; eauto. + exploit me_flat; eauto. apply ATree.elements_complete; eauto. intros [P Q]. subst delta. eapply free_blocks_of_env_perm_1 with (m := m); eauto. rewrite <- comp_env_preserved. omega. Qed. @@ -1708,7 +1708,7 @@ Proof. intros. rewrite <- H7. eapply free_list_load; eauto. unfold blocks_of_env; intros. exploit list_in_map_inv; eauto. intros [[id [b1 ty]] [P Q]]. simpl in P. inv P. - exploit me_range; eauto. eapply PTree.elements_complete; eauto. xomega. + exploit me_range; eauto. eapply ATree.elements_complete; eauto. xomega. rewrite (free_list_nextblock _ _ _ H3). inv H; xomega. rewrite (free_list_nextblock _ _ _ H4). inv H; xomega. Qed. @@ -2191,7 +2191,7 @@ Proof. apply val_inject_list_incr with j'; eauto. eexact B. eexact C. intros. apply (create_undef_temps_lifted id f). auto. - intros. destruct (create_undef_temps (fn_temps f))!id as [v|] eqn:?; auto. + intros. destruct (create_undef_temps (fn_temps f))$id as [v|] eqn:?; auto. exploit create_undef_temps_inv; eauto. intros [P Q]. elim (l id id); auto. intros [tel [tm1 [P [Q [R [S T]]]]]]. change (cenv_for_gen (addr_taken_stmt (fn_body f)) (fn_params f ++ fn_vars f)) @@ -2211,7 +2211,7 @@ Proof. intros. transitivity (Mem.load chunk m0 b 0). eapply bind_parameters_load; eauto. intros. exploit alloc_variables_range. eexact H1. eauto. - unfold empty_env. rewrite PTree.gempty. intros [?|?]. congruence. + unfold empty_env. rewrite ATree.gempty. intros [?|?]. congruence. red; intros; subst b'. xomega. eapply alloc_variables_load; eauto. apply compat_cenv_for. diff --git a/common/AST.v b/common/AST.v index 145f49190a..160837ec08 100644 --- a/common/AST.v +++ b/common/AST.v @@ -18,18 +18,21 @@ Require Import String. Require Import Coqlib Maps Errors Integers Floats. +Require Export Symbols. Require Archi. Set Implicit Arguments. (** * Syntactic elements *) -(** Identifiers (names of local variables, of global symbols and functions, - etc) are represented by the type [positive] of positive integers. *) +(** The type [ident] provides an efficient representation for the + names of local variables, of global symbols and functions, etc. *) -Definition ident := positive. +Definition ident := Ident.t. -Definition ident_eq := peq. +Definition ident_eq := Ident.eq. + +Notation "# s" := (Ident.of_string s) (at level 1, format "'#' s"). (** The intermediate languages are weakly typed, using the following types: *) @@ -275,8 +278,8 @@ Definition prog_defs_names (F V: Type) (p: program F V) : list ident := (** The "definition map" of a program maps names of globals to their definitions. If several definitions have the same name, the one appearing last in [p.(prog_defs)] wins. *) -Definition prog_defmap (F V: Type) (p: program F V) : PTree.t (globdef F V) := - PTree_Properties.of_list p.(prog_defs). +Definition prog_defmap (F V: Type) (p: program F V) : ATree.t (globdef F V) := + ATree_Properties.of_list p.(prog_defs). Section DEFMAP. @@ -284,33 +287,33 @@ Variables F V: Type. Variable p: program F V. Lemma in_prog_defmap: - forall id g, (prog_defmap p)!id = Some g -> In (id, g) (prog_defs p). + forall id g, (prog_defmap p)$id = Some g -> In (id, g) (prog_defs p). Proof. - apply PTree_Properties.in_of_list. + apply ATree_Properties.in_of_list. Qed. Lemma prog_defmap_dom: - forall id, In id (prog_defs_names p) -> exists g, (prog_defmap p)!id = Some g. + forall id, In id (prog_defs_names p) -> exists g, (prog_defmap p)$id = Some g. Proof. - apply PTree_Properties.of_list_dom. + apply ATree_Properties.of_list_dom. Qed. Lemma prog_defmap_unique: forall defs1 id g defs2, prog_defs p = defs1 ++ (id, g) :: defs2 -> ~In id (map fst defs2) -> - (prog_defmap p)!id = Some g. + (prog_defmap p)$id = Some g. Proof. - unfold prog_defmap; intros. rewrite H. apply PTree_Properties.of_list_unique; auto. + unfold prog_defmap; intros. rewrite H. apply ATree_Properties.of_list_unique; auto. Qed. Lemma prog_defmap_norepet: forall id g, list_norepet (prog_defs_names p) -> In (id, g) (prog_defs p) -> - (prog_defmap p)!id = Some g. + (prog_defmap p)$id = Some g. Proof. - apply PTree_Properties.of_list_norepet. + apply ATree_Properties.of_list_norepet. Qed. End DEFMAP. @@ -519,7 +522,7 @@ Definition ef_reloads (ef: external_function) : bool := Definition external_function_eq: forall (ef1 ef2: external_function), {ef1=ef2} + {ef1<>ef2}. Proof. - generalize ident_eq string_dec signature_eq chunk_eq typ_eq list_eq_dec zeq Int.eq_dec; intros. + generalize ident_eq string_dec signature_eq chunk_eq typ_eq list_eq_dec zeq Int.eq_dec peq; intros. decide equality. Defined. Global Opaque external_function_eq. diff --git a/common/Errors.v b/common/Errors.v index 28933313d5..ecb4940023 100644 --- a/common/Errors.v +++ b/common/Errors.v @@ -17,6 +17,7 @@ Require Import String. Require Import Coqlib. +Require Import Symbols. Close Scope string_scope. @@ -30,7 +31,7 @@ Set Implicit Arguments. Inductive errcode: Type := | MSG: string -> errcode - | CTX: positive -> errcode (* a top-level identifier *) + | CTX: Ident.t -> errcode (* a top-level identifier *) | POS: positive -> errcode. (* a positive integer, e.g. a PC *) Definition errmsg: Type := list errcode. diff --git a/common/Globalenvs.v b/common/Globalenvs.v index d37fbd4637..adca12719e 100644 --- a/common/Globalenvs.v +++ b/common/Globalenvs.v @@ -146,13 +146,13 @@ Variable V: Type. (**r The type of information attached to variables *) Record t: Type := mkgenv { genv_public: list ident; (**r which symbol names are public *) - genv_symb: PTree.t block; (**r mapping symbol -> block *) + genv_symb: ATree.t block; (**r mapping symbol -> block *) genv_defs: PTree.t (globdef F V); (**r mapping block -> definition *) genv_next: block; (**r next symbol pointer *) - genv_symb_range: forall id b, PTree.get id genv_symb = Some b -> Plt b genv_next; + genv_symb_range: forall id b, ATree.get id genv_symb = Some b -> Plt b genv_next; genv_defs_range: forall b g, PTree.get b genv_defs = Some g -> Plt b genv_next; genv_vars_inj: forall id1 id2 b, - PTree.get id1 genv_symb = Some b -> PTree.get id2 genv_symb = Some b -> id1 = id2 + ATree.get id1 genv_symb = Some b -> ATree.get id2 genv_symb = Some b -> id1 = id2 }. (** ** Lookup functions *) @@ -160,7 +160,7 @@ Record t: Type := mkgenv { (** [find_symbol ge id] returns the block associated with the given name, if any *) Definition find_symbol (ge: t) (id: ident) : option block := - PTree.get id ge.(genv_symb). + ATree.get id ge.(genv_symb). (** [symbol_address ge id ofs] returns a pointer into the block associated with [id], at byte offset [ofs]. [Vundef] is returned if no block is associated @@ -203,7 +203,7 @@ Definition find_funct (ge: t) (v: val) : option F := (** [invert_symbol ge b] returns the name associated with the given block, if any *) Definition invert_symbol (ge: t) (b: block) : option ident := - PTree.fold + ATree.fold (fun res id b' => if eq_block b b' then Some id else res) ge.(genv_symb) None. @@ -227,13 +227,13 @@ Definition block_is_volatile (ge: t) (b: block) : bool := Program Definition add_global (ge: t) (idg: ident * globdef F V) : t := @mkgenv ge.(genv_public) - (PTree.set idg#1 ge.(genv_next) ge.(genv_symb)) + (ATree.set idg#1 ge.(genv_next) ge.(genv_symb)) (PTree.set ge.(genv_next) idg#2 ge.(genv_defs)) (Pos.succ ge.(genv_next)) _ _ _. Next Obligation. destruct ge; simpl in *. - rewrite PTree.gsspec in H. destruct (peq id i). inv H. apply Plt_succ. + rewrite ATree.gsspec in H. destruct (ATree.elt_eq id i). inv H. apply Plt_succ. apply Plt_trans_succ; eauto. Qed. Next Obligation. @@ -244,8 +244,8 @@ Next Obligation. Qed. Next Obligation. destruct ge; simpl in *. - rewrite PTree.gsspec in H. rewrite PTree.gsspec in H0. - destruct (peq id1 i); destruct (peq id2 i). + rewrite ATree.gsspec in H. rewrite ATree.gsspec in H0. + destruct (ATree.elt_eq id1 i); destruct (ATree.elt_eq id2 i). congruence. inv H. eelim Plt_strict. eapply genv_symb_range0; eauto. inv H0. eelim Plt_strict. eapply genv_symb_range0; eauto. @@ -263,15 +263,15 @@ Proof. Qed. Program Definition empty_genv (pub: list ident): t := - @mkgenv pub (PTree.empty _) (PTree.empty _) 1%positive _ _ _. + @mkgenv pub (ATree.empty _) (PTree.empty _) 1%positive _ _ _. Next Obligation. - rewrite PTree.gempty in H. discriminate. + rewrite ATree.gempty in H. discriminate. Qed. Next Obligation. rewrite PTree.gempty in H. discriminate. Qed. Next Obligation. - rewrite PTree.gempty in H. discriminate. + rewrite ATree.gempty in H. discriminate. Qed. Definition globalenv (p: program F V) := @@ -418,18 +418,18 @@ Qed. Theorem find_def_symbol: forall p id g, - (prog_defmap p)!id = Some g <-> exists b, find_symbol (globalenv p) id = Some b /\ find_def (globalenv p) b = Some g. + (prog_defmap p)$id = Some g <-> exists b, find_symbol (globalenv p) id = Some b /\ find_def (globalenv p) b = Some g. Proof. intros. - set (P := fun m ge => m!id = Some g <-> exists b, find_symbol ge id = Some b /\ find_def ge b = Some g). + set (P := fun m ge => m$id = Some g <-> exists b, find_symbol ge id = Some b /\ find_def ge b = Some g). assert (REC: forall l m ge, P m ge -> - P (fold_left (fun m idg => PTree.set idg#1 idg#2 m) l m) + P (fold_left (fun m idg => ATree.set idg#1 idg#2 m) l m) (add_globals ge l)). { induction l as [ | [id1 g1] l]; intros; simpl. - auto. - apply IHl. unfold P, add_global, find_symbol, find_def; simpl. - rewrite ! PTree.gsspec. destruct (peq id id1). + rewrite ! ATree.gsspec. destruct (ATree.elt_eq id id1). + subst id1. split; intros. inv H0. exists (genv_next ge); split; auto. apply PTree.gss. destruct H0 as (b & A & B). inv A. rewrite PTree.gss in B. auto. @@ -440,7 +440,7 @@ Proof. apply Plt_ne. eapply genv_symb_range; eauto. } apply REC. unfold P, find_symbol, find_def; simpl. - rewrite ! PTree.gempty. split. + rewrite ! ATree.gempty. split. congruence. intros (b & A & B); congruence. Qed. @@ -452,11 +452,11 @@ Theorem find_symbol_exists: Proof. intros. unfold globalenv. eapply add_globals_ensures; eauto. (* preserves *) - intros. unfold find_symbol; simpl. rewrite PTree.gsspec. destruct (peq id id0). + intros. unfold find_symbol; simpl. rewrite ATree.gsspec. destruct (ATree.elt_eq id id0). econstructor; eauto. auto. (* ensures *) - intros. unfold find_symbol; simpl. rewrite PTree.gss. econstructor; eauto. + intros. unfold find_symbol; simpl. rewrite ATree.gss. econstructor; eauto. Qed. Theorem find_symbol_inversion : forall p x b, @@ -465,11 +465,11 @@ Theorem find_symbol_inversion : forall p x b, Proof. intros until b; unfold globalenv. eapply add_globals_preserves. (* preserves *) - unfold find_symbol; simpl; intros. rewrite PTree.gsspec in H1. - destruct (peq x id). subst x. change id with (fst (id, g)). apply List.in_map; auto. + unfold find_symbol; simpl; intros. rewrite ATree.gsspec in H1. + destruct (ATree.elt_eq x id). subst x. change id with (fst (id, g)). apply List.in_map; auto. auto. (* base *) - unfold find_symbol; simpl; intros. rewrite PTree.gempty in H. discriminate. + unfold find_symbol; simpl; intros. rewrite ATree.gempty in H. discriminate. Qed. Theorem find_def_inversion: @@ -538,12 +538,12 @@ Theorem invert_find_symbol: invert_symbol ge b = Some id -> find_symbol ge id = Some b. Proof. intros until b; unfold find_symbol, invert_symbol. - apply PTree_Properties.fold_rec. + apply ATree_Properties.fold_rec. intros. rewrite H in H0; auto. congruence. - intros. destruct (eq_block b v). inv H2. apply PTree.gss. - rewrite PTree.gsspec. destruct (peq id k). - subst. assert (m!k = Some b) by auto. congruence. + intros. destruct (eq_block b v). inv H2. apply ATree.gss. + rewrite ATree.gsspec. destruct (ATree.elt_eq id k). + subst. assert (m$k = Some b) by auto. congruence. auto. Qed. @@ -554,11 +554,11 @@ Proof. intros until b. assert (find_symbol ge id = Some b -> exists id', invert_symbol ge b = Some id'). unfold find_symbol, invert_symbol. - apply PTree_Properties.fold_rec. + apply ATree_Properties.fold_rec. intros. rewrite H in H0; auto. - rewrite PTree.gempty; congruence. + rewrite ATree.gempty; congruence. intros. destruct (eq_block b v). exists k; auto. - rewrite PTree.gsspec in H2. destruct (peq id k). + rewrite ATree.gsspec in H2. destruct (ATree.elt_eq id k). inv H2. congruence. auto. intros; exploit H; eauto. intros [id' A]. @@ -611,8 +611,8 @@ Definition to_senv (ge: t) : Senv.t := (block_is_volatile ge) ge.(genv_next) ge.(genv_vars_inj) - (invert_find_symbol ge) - (find_invert_symbol ge) + (@invert_find_symbol ge) + (@find_invert_symbol ge) (public_symbol_exists ge) ge.(genv_symb_range) (block_is_volatile_below ge). @@ -1672,7 +1672,7 @@ Record match_genvs (ge1: t A V) (ge2: t B W): Prop := { mge_next: genv_next ge2 = genv_next ge1; mge_symb: - forall id, PTree.get id (genv_symb ge2) = PTree.get id (genv_symb ge1); + forall id, ATree.get id (genv_symb ge2) = ATree.get id (genv_symb ge1); mge_defs: forall b, option_rel R (PTree.get b (genv_defs ge1)) (PTree.get b (genv_defs ge2)) }. @@ -1685,7 +1685,7 @@ Lemma add_global_match: Proof. intros. destruct H. constructor; simpl; intros. - congruence. -- rewrite mge_next0, ! PTree.gsspec. destruct (peq id0 id); auto. +- rewrite mge_next0, ! ATree.gsspec. destruct (ATree.elt_eq id0 id); auto. - rewrite mge_next0, ! PTree.gsspec. destruct (peq b (genv_next ge1)). constructor; auto. auto. diff --git a/common/Linking.v b/common/Linking.v index eaa954621e..805ceb7422 100644 --- a/common/Linking.v +++ b/common/Linking.v @@ -246,11 +246,11 @@ Let dm1 := prog_defmap p1. Let dm2 := prog_defmap p2. Definition link_prog_check (x: ident) (gd1: globdef F V) := - match dm2!x with + match dm2$x with | None => true | Some gd2 => - In_dec peq x p1.(prog_public) - && In_dec peq x p2.(prog_public) + In_dec ident_eq x p1.(prog_public) + && In_dec ident_eq x p2.(prog_public) && match link gd1 gd2 with Some _ => true | None => false end end. @@ -263,10 +263,10 @@ Definition link_prog_merge (o1 o2: option (globdef F V)) := Definition link_prog := if ident_eq p1.(prog_main) p2.(prog_main) - && PTree_Properties.for_all dm1 link_prog_check then + && ATree_Properties.for_all dm1 link_prog_check then Some {| prog_main := p1.(prog_main); prog_public := p1.(prog_public) ++ p2.(prog_public); - prog_defs := PTree.elements (PTree.combine link_prog_merge dm1 dm2) |} + prog_defs := ATree.elements (ATree.combine link_prog_merge dm1 dm2) |} else None. @@ -275,20 +275,20 @@ Lemma link_prog_inv: link_prog = Some p -> p1.(prog_main) = p2.(prog_main) /\ (forall id gd1 gd2, - dm1!id = Some gd1 -> dm2!id = Some gd2 -> + dm1$id = Some gd1 -> dm2$id = Some gd2 -> In id p1.(prog_public) /\ In id p2.(prog_public) /\ exists gd, link gd1 gd2 = Some gd) /\ p = {| prog_main := p1.(prog_main); prog_public := p1.(prog_public) ++ p2.(prog_public); - prog_defs := PTree.elements (PTree.combine link_prog_merge dm1 dm2) |}. + prog_defs := ATree.elements (ATree.combine link_prog_merge dm1 dm2) |}. Proof. unfold link_prog; intros p E. destruct (ident_eq (prog_main p1) (prog_main p2)); try discriminate. - destruct (PTree_Properties.for_all dm1 link_prog_check) eqn:C; inv E. - rewrite PTree_Properties.for_all_correct in C. + destruct (ATree_Properties.for_all dm1 link_prog_check) eqn:C; inv E. + rewrite ATree_Properties.for_all_correct in C. split; auto. split; auto. intros. exploit C; eauto. unfold link_prog_check. rewrite H0. intros. - destruct (in_dec peq id (prog_public p1)); try discriminate. - destruct (in_dec peq id (prog_public p2)); try discriminate. + destruct (in_dec ident_eq id (prog_public p1)); try discriminate. + destruct (in_dec ident_eq id (prog_public p2)); try discriminate. destruct (link gd1 gd2) eqn:L; try discriminate. intuition auto. exists g; auto. Qed. @@ -296,26 +296,26 @@ Qed. Lemma link_prog_succeeds: p1.(prog_main) = p2.(prog_main) -> (forall id gd1 gd2, - dm1!id = Some gd1 -> dm2!id = Some gd2 -> + dm1$id = Some gd1 -> dm2$id = Some gd2 -> In id p1.(prog_public) /\ In id p2.(prog_public) /\ link gd1 gd2 <> None) -> link_prog = Some {| prog_main := p1.(prog_main); prog_public := p1.(prog_public) ++ p2.(prog_public); - prog_defs := PTree.elements (PTree.combine link_prog_merge dm1 dm2) |}. + prog_defs := ATree.elements (ATree.combine link_prog_merge dm1 dm2) |}. Proof. intros. unfold link_prog. unfold proj_sumbool. rewrite H, dec_eq_true. simpl. - replace (PTree_Properties.for_all dm1 link_prog_check) with true; auto. - symmetry. apply PTree_Properties.for_all_correct; intros. rename a into gd1. - unfold link_prog_check. destruct dm2!x as [gd2|] eqn:G2; auto. + replace (ATree_Properties.for_all dm1 link_prog_check) with true; auto. + symmetry. apply ATree_Properties.for_all_correct; intros. rename a into gd1. + unfold link_prog_check. destruct dm2$x as [gd2|] eqn:G2; auto. exploit H0; eauto. intros (P & Q & R). unfold proj_sumbool; rewrite ! pred_dec_true by auto. destruct (link gd1 gd2); auto; discriminate. Qed. Lemma prog_defmap_elements: - forall (m: PTree.t (globdef F V)) pub mn x, - (prog_defmap {| prog_defs := PTree.elements m; prog_public := pub; prog_main := mn |})!x = m!x. + forall (m: ATree.t (globdef F V)) pub mn x, + (prog_defmap {| prog_defs := ATree.elements m; prog_public := pub; prog_main := mn |})$x = m$x. Proof. - intros. unfold prog_defmap; simpl. apply PTree_Properties.of_list_elements. + intros. unfold prog_defmap; simpl. apply ATree_Properties.of_list_elements. Qed. End LINKER_PROG. @@ -326,9 +326,9 @@ Instance Linker_prog (F V: Type) {LF: Linker F} {LV: Linker V} : Linker (program p1.(prog_main) = p2.(prog_main) /\ incl p1.(prog_public) p2.(prog_public) /\ forall id gd1, - (prog_defmap p1)!id = Some gd1 -> + (prog_defmap p1)$id = Some gd1 -> exists gd2, - (prog_defmap p2)!id = Some gd2 + (prog_defmap p2)$id = Some gd2 /\ linkorder gd1 gd2 /\ (~In id p2.(prog_public) -> gd2 = gd1) }. @@ -346,15 +346,15 @@ Proof. - intros. apply link_prog_inv in H. destruct H as (L1 & L2 & L3). subst z; simpl. intuition auto. + red; intros; apply in_app_iff; auto. -+ rewrite prog_defmap_elements, PTree.gcombine, H by auto. - destruct (prog_defmap y)!id as [gd2|] eqn:GD2; simpl. ++ rewrite prog_defmap_elements, ATree.gcombine, H by auto. + destruct (prog_defmap y)$id as [gd2|] eqn:GD2; simpl. * exploit L2; eauto. intros (P & Q & gd & R). exists gd; split. auto. split. apply link_linkorder in R; tauto. rewrite in_app_iff; tauto. * exists gd1; split; auto. split. apply linkorder_refl. auto. + red; intros; apply in_app_iff; auto. -+ rewrite prog_defmap_elements, PTree.gcombine, H by auto. - destruct (prog_defmap x)!id as [gd2|] eqn:GD2; simpl. ++ rewrite prog_defmap_elements, ATree.gcombine, H by auto. + destruct (prog_defmap x)$id as [gd2|] eqn:GD2; simpl. * exploit L2; eauto. intros (P & Q & gd & R). exists gd; split. auto. split. apply link_linkorder in R; tauto. rewrite in_app_iff; tauto. @@ -364,8 +364,8 @@ Defined. Lemma prog_defmap_linkorder: forall {F V: Type} {LF: Linker F} {LV: Linker V} (p1 p2: program F V) id gd1, linkorder p1 p2 -> - (prog_defmap p1)!id = Some gd1 -> - exists gd2, (prog_defmap p2)!id = Some gd2 /\ linkorder gd1 gd2. + (prog_defmap p1)$id = Some gd1 -> + exists gd2, (prog_defmap p2)$id = Some gd2 /\ linkorder gd1 gd2. Proof. intros. destruct H as (A & B & C). exploit C; eauto. intros (gd2 & P & Q & R). exists gd2; auto. @@ -415,9 +415,9 @@ Definition match_program_gen (ctx: C) (p1: program F1 V1) (p2: program F2 V2) : Theorem match_program_defmap: forall ctx p1 p2, match_program_gen ctx p1 p2 -> - forall id, option_rel (match_globdef ctx) (prog_defmap p1)!id (prog_defmap p2)!id. + forall id, option_rel (match_globdef ctx) (prog_defmap p1)$id (prog_defmap p2)$id. Proof. - intros. apply PTree_Properties.of_list_related. apply H. + intros. apply ATree_Properties.of_list_related. apply H. Qed. Lemma match_program_gen_main: @@ -644,8 +644,8 @@ Proof. exploit link_match_globdef. eexact H2. eexact H3. eauto. eauto. eauto. intros (tg & TL & _). intuition congruence. - split; [|split]. -+ rewrite R. apply PTree.elements_canonical_order'. intros id. - rewrite ! PTree.gcombine by auto. ++ rewrite R. apply ATree.elements_canonical_order'. intros id. + rewrite ! ATree.gcombine by auto. generalize (match_program_defmap _ _ _ _ _ H0 id) (match_program_defmap _ _ _ _ _ H1 id). clear R. intros R1 R2; inv R1; inv R2; unfold link_prog_merge. * constructor. diff --git a/driver/Interp.ml b/driver/Interp.ml index 6760e76c73..818a783f63 100644 --- a/driver/Interp.ml +++ b/driver/Interp.ml @@ -95,7 +95,7 @@ let name_of_function prog fn = in find_name prog.Ctypes.prog_defs let invert_local_variable e b = - Maps.PTree.fold + Symbols.ATree.fold (fun res id (b', _) -> if b = b' then Some id else res) e None @@ -129,12 +129,12 @@ let print_state p (prog, ge, s) = (name_of_function prog f) PrintCsyntax.print_expr r | Callstate(fd, args, k, m) -> - PrintCsyntax.print_pointer_hook := print_pointer ge.genv_genv Maps.PTree.empty; + PrintCsyntax.print_pointer_hook := print_pointer ge.genv_genv Symbols.ATree.empty; fprintf p "calling@ @[%s(%a)@]" (name_of_fundef prog fd) print_val_list args | Returnstate(res, k, m) -> - PrintCsyntax.print_pointer_hook := print_pointer ge.genv_genv Maps.PTree.empty; + PrintCsyntax.print_pointer_hook := print_pointer ge.genv_genv Symbols.ATree.empty; fprintf p "returning@ %a" print_val res | Stuckstate -> diff --git a/exportclight/Clightnorm.ml b/exportclight/Clightnorm.ml index 4b01d77770..22e146ae1f 100644 --- a/exportclight/Clightnorm.ml +++ b/exportclight/Clightnorm.ml @@ -33,12 +33,12 @@ open Camlcoq open Ctypes open Clight -let gen_next : AST.ident ref = ref P.one +let gen_next : int ref = ref 1 let gen_trail : (AST.ident * coq_type) list ref = ref [] let gensym ty = - let id = !gen_next in - gen_next := P.succ id; + let id = intern_string (Printf.sprintf "@%d" !gen_next) in + gen_next := !gen_next + 1; gen_trail := (id, ty) :: !gen_trail; id @@ -148,10 +148,7 @@ let next_var curr (v, _) = if P.lt v curr then curr else P.succ v let next_var_list vars start = List.fold_left next_var start vars let norm_function f = - gen_next := next_var_list f.fn_params - (next_var_list f.fn_vars - (next_var_list f.fn_temps - (Camlcoq.first_unused_ident ()))); + gen_next := 1; gen_trail := []; let s' = norm_stmt f.fn_body in let new_temps = !gen_trail in diff --git a/exportclight/ExportClight.ml b/exportclight/ExportClight.ml index 1b1402c36e..011c045099 100644 --- a/exportclight/ExportClight.ml +++ b/exportclight/ExportClight.ml @@ -47,59 +47,8 @@ let print_list fn p l = exception Not_an_identifier -let sanitize s = - let s' = Bytes.create (String.length s) in - for i = 0 to String.length s - 1 do - Bytes.set s' i - (match String.get s i with - | 'A'..'Z' | 'a'..'z' | '0'..'9' | '_' as c -> c - | ' ' | '$' -> '_' - | _ -> raise Not_an_identifier) - done; - Bytes.to_string s' - -let temp_names : (ident, string) Hashtbl.t = Hashtbl.create 17 - let ident p id = - try - let s = Hashtbl.find string_of_atom id in - fprintf p "_%s" (sanitize s) - with Not_found | Not_an_identifier -> - try - let s = Hashtbl.find temp_names id in - fprintf p "%s" s - with Not_found -> - fprintf p "%ld%%positive" (P.to_int32 id) - -let iter_hashtbl_sorted (h: ('a, string) Hashtbl.t) (f: 'a * string -> unit) = - List.iter f - (List.fast_sort (fun (k1, d1) (k2, d2) -> String.compare d1 d2) - (Hashtbl.fold (fun k d accu -> (k, d) :: accu) h [])) - -let define_idents p = - iter_hashtbl_sorted - string_of_atom - (fun (id, name) -> - try - fprintf p "Definition _%s : ident := %ld%%positive.@ " - (sanitize name) (P.to_int32 id) - with Not_an_identifier -> - ()); - iter_hashtbl_sorted - temp_names - (fun (id, name) -> - fprintf p "Definition %s : ident := %ld%%positive.@ " - name (P.to_int32 id)); - fprintf p "@ " - -let name_temporary t = - let t1 = P.to_int t and t0 = P.to_int (first_unused_ident ()) in - if t1 >= t0 && not (Hashtbl.mem temp_names t) - then Hashtbl.add temp_names t (sprintf "_t'%d" (t1 - t0 + 1)) - -let name_opt_temporary = function - | None -> () - | Some id -> name_temporary id + fprintf p "#\"%s\"" (extern_atom id) (* Numbers *) @@ -263,7 +212,7 @@ let external_function p = function assertions := (camlstring_of_coqstring text, [targ]) :: !assertions; fprintf p "(EF_annot_val %a %a)" coqstring text asttype targ | EF_debug(kind, text, targs) -> - fprintf p "(EF_debug %ld%%positive %ld%%positive %a)" (P.to_int32 kind) (P.to_int32 text) (print_list asttype) targs + fprintf p "(EF_debug %ld%%positive %a %a)" (P.to_int32 kind) ident text (print_list asttype) targs | EF_inline_asm(text, sg, clob) -> fprintf p "@[(EF_inline_asm %a@ %a@ %a)@]" coqstring text @@ -490,67 +439,11 @@ From compcert Require Import Coqlib Integers Floats AST Ctypes Cop Clight Clight Local Open Scope Z_scope.\n\ \n" -(* Naming the compiler-generated temporaries occurring in the program *) - -let rec name_expr = function - | Evar(id, t) -> () - | Etempvar(id, t) -> name_temporary id - | Ederef(a1, t) -> name_expr a1 - | Efield(a1, f, t) -> name_expr a1 - | Econst_int(n, t) -> () - | Econst_float(n, t) -> () - | Econst_long(n, t) -> () - | Econst_single(n, t) -> () - | Eunop(op, a1, t) -> name_expr a1 - | Eaddrof(a1, t) -> name_expr a1 - | Ebinop(op, a1, a2, t) -> name_expr a1; name_expr a2 - | Ecast(a1, t) -> name_expr a1 - | Esizeof(t1, t) -> () - | Ealignof(t1, t) -> () - -let rec name_stmt = function - | Sskip -> () - | Sassign(e1, e2) -> name_expr e1; name_expr e2 - | Sset(id, e2) -> name_temporary id; name_expr e2 - | Scall(optid, e1, el) -> - name_opt_temporary optid; name_expr e1; List.iter name_expr el - | Sbuiltin(optid, ef, tyl, el) -> - name_opt_temporary optid; List.iter name_expr el - | Ssequence(s1, s2) -> name_stmt s1; name_stmt s2 - | Sifthenelse(e, s1, s2) -> name_expr e; name_stmt s1; name_stmt s2 - | Sloop(s1, s2) -> name_stmt s1; name_stmt s2 - | Sbreak -> () - | Scontinue -> () - | Sswitch(e, cases) -> name_expr e; name_lblstmts cases - | Sreturn (Some e) -> name_expr e - | Sreturn None -> () - | Slabel(lbl, s1) -> name_stmt s1 - | Sgoto lbl -> () - -and name_lblstmts = function - | LSnil -> () - | LScons(lbl, s, ls) -> name_stmt s; name_lblstmts ls - -let name_function f = - List.iter (fun (id, ty) -> name_temporary id) f.fn_temps; - name_stmt f.fn_body - -let name_globdef (id, g) = - match g with - | Gfun(Ctypes.Internal f) -> name_function f - | _ -> () - -let name_program p = - List.iter name_globdef p.Ctypes.prog_defs - (* All together *) let print_program p prog = - Hashtbl.clear temp_names; - name_program prog; fprintf p "@["; fprintf p "%s" prologue; - define_idents p; List.iter (print_globdef p) prog.Ctypes.prog_defs; fprintf p "Definition composites : list composite_definition :=@ "; print_list print_composite_definition p prog.prog_types; diff --git a/extraction/extraction.v b/extraction/extraction.v index a47a72373a..c283bedebd 100644 --- a/extraction/extraction.v +++ b/extraction/extraction.v @@ -47,6 +47,10 @@ Extract Inlined Constant Coqlib.proj_sumbool => "(fun x -> x)". Extract Inlined Constant Datatypes.fst => "fst". Extract Inlined Constant Datatypes.snd => "snd". +(* Symbols *) +Extract Constant Symbols.Ident.of_string => "FastIdent.of_coqstring". +Extract Constant Symbols.Ident.to_string => "FastIdent.to_coqstring". + (* Decidable *) Extraction Inline DecidableClass.Decidable_witness DecidableClass.decide @@ -90,7 +94,6 @@ Extract Constant Allocation.regalloc => "Regalloc.regalloc". Extract Constant Linearize.enumerate_aux => "Linearizeaux.enumerate_aux". (* SimplExpr *) -Extract Constant SimplExpr.first_unused_ident => "Camlcoq.first_unused_ident". Extraction Inline SimplExpr.ret SimplExpr.error SimplExpr.bind SimplExpr.bind2. (* Compopts *) diff --git a/lib/Camlcoq.ml b/lib/Camlcoq.ml index 5c25796ec7..d25dbbbe7a 100644 --- a/lib/Camlcoq.ml +++ b/lib/Camlcoq.ml @@ -275,31 +275,6 @@ let coqint_of_camlint64 : int64 -> Integers.Int64.int = Z.of_uint64 let camlint64_of_ptrofs : Integers.Ptrofs.int -> int64 = fun x -> Z.to_int64 (Integers.Ptrofs.signed x) -(* Atoms (positive integers representing strings) *) - -type atom = positive - -let atom_of_string = (Hashtbl.create 17 : (string, atom) Hashtbl.t) -let string_of_atom = (Hashtbl.create 17 : (atom, string) Hashtbl.t) -let next_atom = ref Coq_xH - -let intern_string s = - try - Hashtbl.find atom_of_string s - with Not_found -> - let a = !next_atom in - next_atom := Pos.succ !next_atom; - Hashtbl.add atom_of_string s a; - Hashtbl.add string_of_atom a s; - a -let extern_atom a = - try - Hashtbl.find string_of_atom a - with Not_found -> - Printf.sprintf "$%d" (P.to_int a) - -let first_unused_ident () = !next_atom - (* Strings *) let camlstring_of_coqstring (s: char list) = @@ -324,6 +299,15 @@ let coqstring_uppercase_ascii_of_camlstring s = cstring (d :: accu) (pos - 1) in cstring [] (String.length s - 1) +(* Atoms (positive integers representing strings) *) + +type atom = Symbols.Ident.t + +let intern_string s = + Symbols.Ident.of_string (coqstring_of_camlstring s) +let extern_atom a = + camlstring_of_coqstring (Symbols.Ident.to_string a) + (* Floats *) let coqfloat_of_camlfloat f = diff --git a/lib/FastIdent.ml b/lib/FastIdent.ml new file mode 100644 index 0000000000..2300588a88 --- /dev/null +++ b/lib/FastIdent.ml @@ -0,0 +1,31 @@ +(* The functions below are used in Symbols.v to implement a fast way to + refer to global identifiers, where they are axiomatized as a bijective + mapping. The bijection is built incrementally, and of_coqstring may + insert a new entry in the hashtable, but cruicially the statefulness + is not observable from the outside. As far as Coq knows, the mapping + existed all along and we simply observed it for the first time. + + The functions below should not be used directly from ML code. Instead, + the Camlcoq module provides wrappers to access them through the Ident + module defined in Symbols.v. *) + +open BinNums +open BinPos + +let atom_of_coqstring = (Hashtbl.create 17 : (char list, positive) Hashtbl.t) +let coqstring_of_atom = (Hashtbl.create 17 : (positive, char list) Hashtbl.t) +let next_atom = ref Coq_xH + +let of_coqstring s = + try + Hashtbl.find atom_of_coqstring s + with Not_found -> + let a = !next_atom in + next_atom := Pos.succ !next_atom; + Hashtbl.add atom_of_coqstring s a; + Hashtbl.add coqstring_of_atom a s; + a + +let to_coqstring a = + Hashtbl.find coqstring_of_atom a + diff --git a/lib/FastIdent.mli b/lib/FastIdent.mli new file mode 100644 index 0000000000..207f381a46 --- /dev/null +++ b/lib/FastIdent.mli @@ -0,0 +1,4 @@ +open Camlcoq + +val of_coqstring : char list -> P.t +val to_coqstring : P.t -> char list diff --git a/lib/Maps.v b/lib/Maps.v index cfb866baa8..ea12d5c11a 100644 --- a/lib/Maps.v +++ b/lib/Maps.v @@ -128,6 +128,19 @@ Module Type TREE. Axiom elements_keys_norepet: forall (A: Type) (m: t A), list_norepet (List.map (@fst elt A) (elements m)). + Axiom elements_canonical_order: + forall (A B: Type) (R: A -> B -> Prop) (m: t A) (n: t B), + (forall i x, get i m = Some x -> exists y, get i n = Some y /\ R x y) -> + (forall i y, get i n = Some y -> exists x, get i m = Some x /\ R x y) -> + list_forall2 + (fun i_x i_y => fst i_x = fst i_y /\ R (snd i_x) (snd i_y)) + (elements m) (elements n). + Axiom elements_canonical_order': + forall (A B: Type) (R: A -> B -> Prop) (m: t A) (n: t B), + (forall i, option_rel R (get i m) (get i n)) -> + list_forall2 + (fun i_x i_y => fst i_x = fst i_y /\ R (snd i_x) (snd i_y)) + (elements m) (elements n). Axiom elements_extensional: forall (A: Type) (m n: t A), (forall i, get i m = get i n) -> diff --git a/lib/Symbols.v b/lib/Symbols.v new file mode 100644 index 0000000000..a5a21217b5 --- /dev/null +++ b/lib/Symbols.v @@ -0,0 +1,110 @@ +Require Import String. +Require Import FSets. +Require Import Coqlib. +Require Import Maps. +Require Import Ordered. + +(** The mapping between the external representation of identifiers as + [string]s and their internal representation as [ident]s is axiomatized + as a bijection, and constructed incrementally by Camlcoq.ml. + + In the future we may want to have an alternative implementation the + [IDENT] interface entirely in Coq, which the user could choose at + [./configure] time depending on their needs. Such an implementation + would have to forgo string interning and use less efficient + primitives, but would compute within Coq and avoid relying on + trusted ML code. *) + +Module Type IDENT. + Parameter t: Type. + Parameter eq: forall i j: t, {i = j} + {i <> j}. + Parameter of_string: string -> t. + Parameter to_string: t -> string. + Parameter lt: t -> t -> Prop. + Parameter compare: t -> t -> comparison. + + Axiom of_to_string: forall i, of_string (to_string i) = i. + Axiom to_of_string: forall s, to_string (of_string s) = s. + + Axiom lt_trans: forall i j k, lt i j -> lt j k -> lt i k. + Axiom lt_not_eq: forall i j, lt i j -> i <> j. + + Axiom compare_spec: + forall i j, CompareSpec (i = j) (lt i j) (lt j i) (compare i j). + + Declare Module Tree : TREE + with Definition elt := t + with Definition elt_eq := eq. + + (** This is a temporary workaround to enable the hack in Initializers.v, + whereby identifiers are stored in the [block] components of pointers. + We can remove it if we merge PR #220 or something such. *) + Parameter to_positive: t -> positive. + Parameter of_positive: positive -> t. + Axiom of_to_positive: forall p, of_positive (to_positive p) = p. +End IDENT. + +Module Ident : IDENT. + Definition t := positive. + Definition eq := peq. + Definition lt := Pos.lt. + Definition compare := Pos.compare. + Definition lt_trans := Pos.lt_trans. + Definition compare_spec := Pos.compare_spec. + + Lemma lt_not_eq i j: + lt i j -> i <> j. + Proof. + intros H Hij. subst. + eapply Pos.lt_irrefl; eauto. + Qed. + + Module Tree := PTree. + + Parameter of_string: string -> t. + Parameter to_string: t -> string. + + Axiom of_to_string: forall i, of_string (to_string i) = i. + Axiom to_of_string: forall s, to_string (of_string s) = s. + + Definition to_positive i: t := i. + Definition of_positive i: t := i. + Definition of_to_positive i: to_positive (of_positive i) = i := eq_refl. +End Ident. + +Module OrderedIdent <: OrderedType. + Definition t := Ident.t. + Definition eq := @eq Ident.t. + Definition lt := Ident.lt. + Definition eq_dec := Ident.eq. + Definition eq_refl := @eq_refl Ident.t. + Definition eq_trans := @eq_trans Ident.t. + Definition eq_sym := @eq_sym Ident.t. + Definition lt_trans := Ident.lt_trans. + Definition lt_not_eq := Ident.lt_not_eq. + + Program Definition compare x y: Compare lt eq x y := + match Ident.compare x y with Lt => LT _ | Eq => EQ _ | Gt => GT _ end. + Next Obligation. + destruct (Ident.compare_spec x y); try discriminate; eauto. + Qed. + Next Obligation. + destruct (Ident.compare_spec x y); try discriminate; eauto. + Qed. + Next Obligation. + destruct (Ident.compare_spec x y); try discriminate; eauto. + Qed. +End OrderedIdent. + +Lemma ident_of_string_inj s t: + Ident.of_string s = Ident.of_string t -> s = t. +Proof. + intros Hst. + rewrite <- (Ident.to_of_string s). + rewrite <- (Ident.to_of_string t). + congruence. +Qed. + +Module ATree := Ident.Tree. +Module ATree_Properties := Tree_Properties ATree. +Notation "a $ b" := (ATree.get b a) (at level 1). diff --git a/powerpc/AsmToJSON.ml b/powerpc/AsmToJSON.ml index 7c4c8f8a14..78d22dd92e 100644 --- a/powerpc/AsmToJSON.ml +++ b/powerpc/AsmToJSON.ml @@ -70,7 +70,7 @@ type instruction_arg = | ALabel of positive | Float32 of Floats.float32 | Float64 of Floats.float - | Atom of positive + | Atom of ident | String of string let pp_arg pp = function diff --git a/x86/TargetPrinter.ml b/x86/TargetPrinter.ml index 3171847cf6..14ce1cad5a 100644 --- a/x86/TargetPrinter.ml +++ b/x86/TargetPrinter.ml @@ -89,17 +89,17 @@ let rec log2 n = module type SYSTEM = sig val raw_symbol: out_channel -> string -> unit - val symbol: out_channel -> P.t -> unit + val symbol: out_channel -> ident -> unit val label: out_channel -> int -> unit val name_of_section: section_name -> string val stack_alignment: int val print_align: out_channel -> int -> unit val print_mov_rs: out_channel -> ireg -> ident -> unit - val print_fun_info: out_channel -> P.t -> unit - val print_var_info: out_channel -> P.t -> unit + val print_fun_info: out_channel -> ident -> unit + val print_var_info: out_channel -> ident -> unit val print_epilogue: out_channel -> unit - val print_comm_decl: out_channel -> P.t -> Z.t -> int -> unit - val print_lcomm_decl: out_channel -> P.t -> Z.t -> int -> unit + val print_comm_decl: out_channel -> ident -> Z.t -> int -> unit + val print_lcomm_decl: out_channel -> ident -> Z.t -> int -> unit end (* Printer functions for ELF *)