(* *********************************************************************)
(*                                                                     *)
(*              The Compcert verified compiler                         *)
(*                                                                     *)
(*          Xavier Leroy, INRIA Paris-Rocquencourt                     *)
(*                                                                     *)
(*  Copyright Institut National de Recherche en Informatique et en     *)
(*  Automatique.  All rights reserved.  This file is distributed       *)
(*  under the terms of the INRIA Non-Commercial License Agreement.     *)
(*                                                                     *)
(* *********************************************************************)

(** Locations are a refinement of RTL pseudo-registers, used to reflect
  the results of register allocation (file [Allocation]). *)

Require Import OrderedType.
Require Import Coqlib.
Require Import Maps.
Require Import Ordered.
Require Import AST.
Require Import Values.
Require Export Machregs.
(** * Representation of locations *) (** A location is either a processor register or (an abstract designation of) a slot in the activation record of the current function. *) (** ** Processor registers *) (** Processor registers usable for register allocation are defined in module [Machregs]. *) (** ** Slots in activation records *) (** A slot in an activation record is designated abstractly by a kind, a type and an integer offset. Three kinds are considered: - [Local]: these are the slots used by register allocation for pseudo-registers that cannot be assigned a hardware register. - [Incoming]: used to store the parameters of the current function that cannot reside in hardware registers, as determined by the calling conventions. - [Outgoing]: used to store arguments to called functions that cannot reside in hardware registers, as determined by the calling conventions. *) Inductive slot: Type := | Local | Incoming | Outgoing.
Proof.
intros p.
pattern sumbool.
destruct p.
destruct q.
eauto.
right.
congruence.
right.
congruence.
destruct q.
right.
congruence.
eauto.
right.
congruence.
destruct q.
right.
congruence.
right.
congruence.
eauto.
Qed.
Proof.
decide equality.
Defined.
Open Scope Z_scope.
Definition typesize (ty: typ) : Z := match ty with | Tint => 1 | Tlong => 2 | Tfloat => 2 | Tsingle => 1 | Tany32 => 1 | Tany64 => 2 end.
Proof.
intros.
unfold typesize.
compute.
destruct ty.
eauto.
eauto.
eauto.
eauto.
eauto.
eauto.
Qed.
Proof.
destruct ty.
{
compute.
auto.
}
{
compute.
auto.
}
{
compute.
auto.
}
{
compute.
auto.
}
{
compute.
auto.
}
{
compute.
auto.
}
Qed.
Definition typealign (ty: typ) : Z := match ty with | Tint => 1 | Tlong => 2 | Tfloat => 1 | Tsingle => 1 | Tany32 => 1 | Tany64 => 1 end.
Proof.
intros.
unfold typealign.
destruct ty.
omega.
omega.
omega.
omega.
omega.
omega.
Qed.
Proof.
destruct ty.
{
compute.
auto.
}
{
compute.
auto.
}
{
compute.
auto.
}
{
compute.
auto.
}
{
compute.
auto.
}
{
compute.
auto.
}
Qed.
Proof.
intros.
destruct ty.
simpl.
destruct Zpos.
intuition.
destruct Zpos.
intuition.
intuition.
intuition.
destruct Zneg.
intuition.
intuition.
intuition.
simpl.
intuition.
simpl.
destruct Zpos.
intuition.
destruct Zpos.
intuition.
intuition.
intuition.
destruct Zneg.
intuition.
intuition.
intuition.
simpl.
destruct Zpos.
intuition.
destruct Zpos.
intuition.
intuition.
intuition.
destruct Zneg.
intuition.
intuition.
intuition.
simpl.
destruct Zpos.
intuition.
destruct Zpos.
intuition.
intuition.
intuition.
destruct Zneg.
intuition.
intuition.
intuition.
simpl.
intuition.
Qed.
Proof.
intros.
exists (typesize ty / typealign ty).
destruct ty.
{
reflexivity.
}
{
reflexivity.
}
{
reflexivity.
}
{
reflexivity.
}
{
reflexivity.
}
{
reflexivity.
}
Qed.
(** ** Locations *) (** Locations are just the disjoint union of machine registers and activation record slots. *) Inductive loc : Type := | R (r: mreg) | S (sl: slot) (pos: Z) (ty: typ).
Module Loc.
Definition type (l: loc) : typ := match l with | R r => mreg_type r | S sl pos ty => ty end.
Proof.
Admitted.
Proof.
decide equality.
{
apply mreg_eq.
}
{
apply typ_eq.
}
{
apply zeq.
}
{
apply slot_eq.
}
Defined.
(** As mentioned previously, two locations can be different (in the sense of the [<>] mathematical disequality), yet denote overlapping memory chunks within the activation record. Given two locations, three cases are possible: - They are equal (in the sense of the [=] equality) - They are different and non-overlapping. - They are different but overlapping. The second case (different and non-overlapping) is characterized by the following [Loc.diff] predicate. *) Definition diff (l1 l2: loc) : Prop := match l1, l2 with | R r1, R r2 => r1 <> r2 | S s1 d1 t1, S s2 d2 t2 => s1 <> s2 \/ d1 + typesize t1 <= d2 \/ d2 + typesize t2 <= d1 | _, _ => True end.
Proof.
intros.
elim l.
eauto.
intros.
elim ty.
simpl.
intuition.
simpl.
intuition.
simpl.
intuition.
simpl.
intuition.
simpl.
intuition.
simpl.
intuition.
Qed.
Proof.
destruct l.
{
unfold diff.
auto.
}
{
unfold diff.
auto.
red.
intros.
destruct H.
{
auto.
}
{
auto.
generalize (typesize_pos ty).
omega.
}
}
Qed.
Proof.
Admitted.
Proof.
unfold not.
intros.
subst l2.
elim (same_not_diff l1 H).
Qed.
Proof.
intros.
destruct l1.
red.
destruct l2.
congruence.
eauto.
simpl in H.
unfold diff.
destruct l2.
eauto.
intuition.
Qed.
Proof.
destruct l1.
{
destruct l2.
{
unfold diff.
auto.
}
{
unfold diff.
auto.
}
}
{
destruct l2.
{
unfold diff.
auto.
}
{
unfold diff.
auto.
intuition.
}
}
Qed.
Proof.
Admitted.
Proof.
intros.
destruct l1.
{
destruct l2.
{
simpl.
destruct (mreg_eq r r0).
{
right.
tauto.
}
{
left.
auto.
}
}
{
simpl.
left.
auto.
}
}
{
destruct l2.
{
simpl.
left.
auto.
}
{
simpl.
destruct (slot_eq sl sl0).
{
destruct (zle (pos + typesize ty) pos0).
{
left.
auto.
}
{
destruct (zle (pos0 + typesize ty0) pos).
{
left.
auto.
}
{
right.
red.
intros [P | [P | P]].
{
congruence.
}
{
omega.
}
{
omega.
}
}
}
}
{
left.
auto.
}
}
}
Defined.
(** We now redefine some standard notions over lists, using the [Loc.diff] predicate instead of standard disequality [<>]. [Loc.notin l ll] holds if the location [l] is different from all locations in the list [ll]. *) Fixpoint notin (l: loc) (ll: list loc) {struct ll} : Prop := match ll with | nil => True | l1 :: ls => diff l l1 /\ notin l ls end.
Proof.
Admitted.
Proof.
induction ll.
{
simpl.
tauto.
}
{
simpl.
rewrite IHll.
intuition.
subst a.
auto.
}
Qed.
Proof.
Admitted.
Proof.
intros.
red.
intros.
rewrite notin_iff in H.
elim (diff_not_eq l l).
{
auto.
}
{
auto.
}
Qed.
Proof.
Admitted.
Proof.
induction ll.
{
simpl.
left.
auto.
}
{
simpl.
destruct (diff_dec l a).
{
destruct IHll.
{
left.
auto.
}
{
right.
tauto.
}
}
{
right.
tauto.
}
}
Defined.
(** [Loc.disjoint l1 l2] is true if the locations in list [l1] are different from all locations in list [l2]. *) Definition disjoint (l1 l2: list loc) : Prop := forall x1 x2, In x1 l1 -> In x2 l2 -> diff x1 x2.
Proof.
intros.
red.
intros x1.
unfold In.
destruct x2.
intuition.
intuition.
Qed.
Proof.
unfold disjoint.
intros.
auto with coqlib.
Qed.
Proof.
intros.
red.
intros.
apply H.
eauto.
simpl.
eauto.
Qed.
Proof.
unfold disjoint.
intros.
auto with coqlib.
Qed.
Proof.
Admitted.
Proof.
unfold disjoint.
intros.
apply diff_sym.
auto.
Qed.
Proof.
Admitted.
Proof.
intros.
rewrite notin_iff in H.
auto.
Qed.
Proof.
Admitted.
Proof.
intros.
red.
intros.
exploit H.
{
eauto.
}
{
eauto.
rewrite notin_iff.
intros.
auto.
}
Qed.
Proof.
Admitted.
Proof.
intros.
rewrite notin_iff.
intros.
red in H.
auto.
Qed.
(** [Loc.norepet ll] holds if the locations in list [ll] are pairwise different. *) Inductive norepet : list loc -> Prop := | norepet_nil: norepet nil | norepet_cons: forall hd tl, notin hd tl -> norepet tl -> norepet (hd :: tl).
Proof.
Admitted.
Proof.
induction ll.
{
left.
constructor.
}
{
destruct (notin_dec a ll).
{
destruct IHll.
{
left.
constructor.
{
auto.
}
{
auto.
}
}
{
right.
red.
intros P.
inv P.
contradiction.
}
}
{
right.
red.
intros P.
inv P.
contradiction.
}
}
Defined.
(** [Loc.no_overlap l1 l2] holds if elements of [l1] never overlap partially with elements of [l2]. *) Definition no_overlap (l1 l2 : list loc) := forall r, In r l1 -> forall s, In s l2 -> r = s \/ Loc.diff r s.
End Loc.
(** * Mappings from locations to values *) (** The [Locmap] module defines mappings from locations to values, used as evaluation environments for the semantics of the [LTL] and [Linear] intermediate languages. *) Set Implicit Arguments.
Module Locmap.
Definition t := loc -> val.
Definition init (x: val) : t := fun (_: loc) => x.
Definition get (l: loc) (m: t) : val := m l.
(** The [set] operation over location mappings reflects the overlapping properties of locations: changing the value of a location [l] invalidates (sets to [Vundef]) the locations that partially overlap with [l]. In other terms, the result of [set l v m] maps location [l] to value [v], locations that overlap with [l] to [Vundef], and locations that are different (and non-overlapping) from [l] to their previous values in [m]. This is apparent in the ``good variables'' properties [Locmap.gss] and [Locmap.gso]. Additionally, the [set] operation also anticipates the fact that abstract stack slots are mapped to concrete memory locations in the [Stacking] phase. Hence, values stored in stack slots are normalized according to the type of the slot. *) Definition set (l: loc) (v: val) (m: t) : t := fun (p: loc) => if Loc.eq l p then match l with R r => v | S sl ofs ty => Val.load_result (chunk_of_type ty) v end else if Loc.diff_dec l p then m p else Vundef.
Proof.
intros.
unfold Val.load_result.
unfold set.
destruct Loc.eq.
eauto.
destruct Loc.diff_dec.
intuition.
intuition.
Qed.
Proof.
intros.
unfold set.
apply dec_eq_true.
Qed.
Proof.
unfold set.
intros.
destruct Loc.eq.
eauto.
destruct Loc.diff_dec.
destruct r.
congruence.
congruence.
congruence.
congruence.
congruence.
congruence.
congruence.
congruence.
congruence.
congruence.
congruence.
congruence.
congruence.
congruence.
congruence.
congruence.
congruence.
congruence.
congruence.
congruence.
congruence.
congruence.
congruence.
congruence.
congruence.
congruence.
congruence.
congruence.
congruence.
congruence.
congruence.
congruence.
destruct v.
eauto.
congruence.
congruence.
congruence.
congruence.
congruence.
Qed.
Proof.
intros.
unfold set.
rewrite dec_eq_true.
auto.
Qed.
Proof.
Admitted.
Proof.
intros.
rewrite gss.
destruct l.
{
auto.
}
{
apply Val.load_result_same.
auto.
}
Qed.
Proof.
Admitted.
Proof.
intros.
unfold set.
destruct (Loc.eq l p).
{
subst p.
elim (Loc.same_not_diff _ H).
}
{
destruct (Loc.diff_dec l p).
{
auto.
}
{
contradiction.
}
}
Qed.
Fixpoint undef (ll: list loc) (m: t) {struct ll} : t := match ll with | nil => m | l1 :: ll' => undef ll' (set l1 Vundef m) end.
Proof.
Admitted.
Proof.
induction ll.
{
simpl.
intros.
auto.
}
{
simpl.
intros.
destruct H.
rewrite IHll.
{
auto.
apply gso.
apply Loc.diff_sym.
auto.
}
{
auto.
}
}
Qed.
Proof.
Admitted.
Proof.
assert (P: forall ll l m, m l = Vundef -> (undef ll m) l = Vundef).
{
induction ll.
{
simpl.
intros.
auto.
}
{
simpl.
intros.
apply IHll.
unfold set.
destruct (Loc.eq a l).
{
destruct a.
{
auto.
}
{
destruct ty.
{
reflexivity.
}
{
reflexivity.
}
{
reflexivity.
}
{
reflexivity.
}
{
reflexivity.
}
{
reflexivity.
}
}
}
{
destruct (Loc.diff_dec a l).
{
auto.
}
{
auto.
}
}
}
}
{
induction ll.
{
simpl.
intros.
contradiction.
}
{
simpl.
intros.
destruct H.
{
apply P.
subst a.
apply gss_typed.
exact I.
}
{
auto.
}
}
}
Qed.
Definition getpair (p: rpair loc) (m: t) : val := match p with | One l => m l | Twolong l1 l2 => Val.longofwords (m l1) (m l2) end.
Definition setpair (p: rpair mreg) (v: val) (m: t) : t := match p with | One r => set (R r) v m | Twolong hi lo => set (R lo) (Val.loword v) (set (R hi) (Val.hiword v) m) end.
Proof.
induction p.
intros.
apply H.
simpl.
eauto.
simpl.
intuition.
rewrite H.
rewrite H.
eauto.
eauto.
eauto.
Qed.
Proof.
intros.
destruct p.
{
simpl.
apply H.
simpl.
auto.
}
{
simpl.
f_equal.
{
apply H.
simpl.
auto.
}
{
apply H.
simpl.
auto.
}
}
Qed.
Proof.
Admitted.
Proof.
intros.
destruct p.
{
simpl in *.
apply gso.
apply Loc.diff_sym.
auto.
}
{
simpl in *.
destruct H.
rewrite ! gso.
{
idtac.
auto.
}
{
apply Loc.diff_sym.
auto .
}
{
apply Loc.diff_sym.
auto .
}
}
Qed.
Fixpoint setres (res: builtin_res mreg) (v: val) (m: t) : t := match res with | BR r => set (R r) v m | BR_none => m | BR_splitlong hi lo => setres lo (Val.loword v) (setres hi (Val.hiword v) m) end.
End Locmap.
(** * Total ordering over locations *) Module IndexedTyp <: INDEXED_TYPE.
Definition t := typ.
Definition index (x: t) := match x with | Tany32 => 1%positive | Tint => 2%positive | Tsingle => 3%positive | Tany64 => 4%positive | Tfloat => 5%positive | Tlong => 6%positive end.
Proof.
unfold index.
intros x.
destruct y.
destruct x.
intuition.
intuition.
congruence.
intuition.
congruence.
intuition.
congruence.
intros.
easy.
intuition.
congruence.
destruct x.
intuition.
congruence.
intros.
eauto.
try discriminate.
intuition.
congruence.
intuition.
congruence.
try discriminate.
destruct x.
intuition.
congruence.
try discriminate.
intros.
eauto.
intuition.
congruence.
intuition.
congruence.
try discriminate.
destruct x.
intuition.
congruence.
intuition.
congruence.
intuition.
congruence.
intuition.
intros.
easy.
intuition.
congruence.
destruct x.
intros.
easy.
intuition.
congruence.
intuition.
congruence.
intros.
easy.
intros.
eauto.
intuition.
congruence.
destruct x.
intuition.
congruence.
try discriminate.
try discriminate.
intuition.
congruence.
intuition.
congruence.
intros.
eauto.
Qed.
Proof.
destruct x.
{
destruct y.
{
simpl.
congruence.
}
{
simpl.
congruence.
}
{
simpl.
congruence.
}
{
simpl.
congruence.
}
{
simpl.
congruence.
}
{
simpl.
congruence.
}
}
{
destruct y.
{
simpl.
congruence.
}
{
simpl.
congruence.
}
{
simpl.
congruence.
}
{
simpl.
congruence.
}
{
simpl.
congruence.
}
{
simpl.
congruence.
}
}
{
destruct y.
{
simpl.
congruence.
}
{
simpl.
congruence.
}
{
simpl.
congruence.
}
{
simpl.
congruence.
}
{
simpl.
congruence.
}
{
simpl.
congruence.
}
}
{
destruct y.
{
simpl.
congruence.
}
{
simpl.
congruence.
}
{
simpl.
congruence.
}
{
simpl.
congruence.
}
{
simpl.
congruence.
}
{
simpl.
congruence.
}
}
{
destruct y.
{
simpl.
congruence.
}
{
simpl.
congruence.
}
{
simpl.
congruence.
}
{
simpl.
congruence.
}
{
simpl.
congruence.
}
{
simpl.
congruence.
}
}
{
destruct y.
{
simpl.
congruence.
}
{
simpl.
congruence.
}
{
simpl.
congruence.
}
{
simpl.
congruence.
}
{
simpl.
congruence.
}
{
simpl.
congruence.
}
}
Qed.
Definition eq := typ_eq.
End IndexedTyp.
Module OrderedTyp := OrderedIndexed(IndexedTyp).
Module IndexedSlot <: INDEXED_TYPE.
Definition t := slot.
Definition index (x: t) := match x with Local => 1%positive | Incoming => 2%positive | Outgoing => 3%positive end.
Proof.
unfold index.
intros x.
destruct y.
destruct x.
intros.
eauto.
intros.
easy.
intros.
easy.
destruct x.
intros.
easy.
intuition.
intuition.
congruence.
destruct x.
intros.
easy.
intuition.
congruence.
intuition.
Qed.
Proof.
destruct x.
{
destruct y.
{
simpl.
congruence.
}
{
simpl.
congruence.
}
{
simpl.
congruence.
}
}
{
destruct y.
{
simpl.
congruence.
}
{
simpl.
congruence.
}
{
simpl.
congruence.
}
}
{
destruct y.
{
simpl.
congruence.
}
{
simpl.
congruence.
}
{
simpl.
congruence.
}
}
Qed.
Definition eq := slot_eq.
End IndexedSlot.
Module OrderedSlot := OrderedIndexed(IndexedSlot).
Module OrderedLoc <: OrderedType.
Definition t := loc.
Definition eq (x y: t) := x = y.
Definition lt (x y: t) := match x, y with | R r1, R r2 => Plt (IndexedMreg.index r1) (IndexedMreg.index r2) | R _, S _ _ _ => True | S _ _ _, R _ => False | S sl1 ofs1 ty1, S sl2 ofs2 ty2 => OrderedSlot.lt sl1 sl2 \/ (sl1 = sl2 /\ (ofs1 < ofs2 \/ (ofs1 = ofs2 /\ OrderedTyp.lt ty1 ty2))) end.
Proof.
unfold eq.
intros.
congruence.
Qed.
Proof (@eq_refl t).
Proof.
unfold eq.
intros.
rewrite H.
eauto.
Qed.
Proof (@eq_sym t).
Proof.
unfold eq.
intros.
rewrite H.
eauto.
Qed.
Proof (@eq_trans t).
Proof.
Admitted.
Proof.
unfold lt.
intros.
destruct x.
{
destruct y.
{
destruct z.
{
try tauto.
eapply Plt_trans.
{
eauto.
}
{
eauto.
}
}
{
try tauto.
}
}
{
destruct z.
{
try tauto.
}
{
try tauto.
}
}
}
{
destruct y.
{
destruct z.
{
try tauto.
}
{
try tauto.
}
}
{
destruct z.
{
try tauto.
}
{
try tauto.
destruct H.
{
destruct H0.
{
left.
eapply OrderedSlot.lt_trans.
{
eauto.
}
{
eauto.
}
}
{
destruct H0.
subst sl0.
auto.
}
}
{
destruct H.
subst sl.
destruct H0.
{
auto.
}
{
destruct H.
right.
split.
{
auto.
}
{
intuition.
right.
split.
{
congruence.
}
{
eapply OrderedTyp.lt_trans.
{
eauto.
}
{
eauto.
}
}
}
}
}
}
}
}
Qed.
Proof.
Admitted.
Proof.
unfold lt.
unfold eq.
intros.
red.
intros.
subst y.
destruct x.
{
eelim Plt_strict.
eauto.
}
{
destruct H.
{
eelim OrderedSlot.lt_not_eq.
{
eauto.
}
{
eauto.
red.
auto.
}
}
{
destruct H.
destruct H0.
{
omega.
}
{
destruct H0.
eelim OrderedTyp.lt_not_eq.
{
eauto.
}
{
eauto.
red.
auto.
}
}
}
}
Qed.
Proof.
Admitted.
Proof.
intros.
destruct x.
{
destruct y.
{
destruct (OrderedPositive.compare (IndexedMreg.index r) (IndexedMreg.index r0)).
{
apply LT.
red.
auto.
}
{
apply EQ.
red.
f_equal.
apply IndexedMreg.index_inj.
auto.
}
{
apply GT.
red.
auto.
}
}
{
apply LT.
red.
auto.
}
}
{
destruct y.
{
apply GT.
red.
auto.
}
{
destruct (OrderedSlot.compare sl sl0).
{
apply LT.
red.
auto.
}
{
destruct (OrderedZ.compare pos pos0).
{
apply LT.
red.
auto.
}
{
destruct (OrderedTyp.compare ty ty0).
{
apply LT.
red.
auto.
}
{
apply EQ.
red.
red in e.
red in e0.
red in e1.
congruence.
}
{
apply GT.
red.
auto.
}
}
{
apply GT.
red.
auto.
}
}
{
apply GT.
red.
auto.
}
}
}
Defined.
Definition eq_dec := Loc.eq.
(** Connection between the ordering defined here and the [Loc.diff] predicate. *) Definition diff_low_bound (l: loc) : loc := match l with | R mr => l | S sl ofs ty => S sl (ofs - 1) Tany64 end.
Definition diff_high_bound (l: loc) : loc := match l with | R mr => l | S sl ofs ty => S sl (ofs + typesize ty - 1) Tlong end.
Proof.
Admitted.
Proof.
intros.
destruct l as [mr | sl ofs ty].
{
destruct l' as [mr' | sl' ofs' ty'].
{
simpl in *.
auto.
assert (IndexedMreg.index mr <> IndexedMreg.index mr').
{
destruct H.
{
apply not_eq_sym.
apply Plt_ne.
auto.
}
{
apply Plt_ne.
auto.
}
}
{
congruence.
}
}
{
simpl in *.
auto.
}
}
{
destruct l' as [mr' | sl' ofs' ty'].
{
simpl in *.
auto.
}
{
simpl in *.
auto.
assert (RANGE: forall ty, 1 <= typesize ty <= 2).
{
intros.
unfold typesize.
destruct ty0.
{
omega.
}
{
omega.
}
{
omega.
}
{
omega.
}
{
omega.
}
{
omega.
}
}
{
destruct H.
{
destruct H.
{
left.
apply not_eq_sym.
apply OrderedSlot.lt_not_eq.
auto.
}
{
destruct H.
right.
destruct H0.
{
right.
generalize (RANGE ty').
omega.
}
{
destruct H0.
assert (ty' = Tint \/ ty' = Tsingle \/ ty' = Tany32).
{
unfold OrderedTyp.lt in H1.
destruct ty'.
{
auto.
}
{
auto.
compute in H1.
congruence.
}
{
auto.
compute in H1.
congruence.
}
{
auto.
}
{
auto.
}
{
auto.
compute in H1.
congruence.
}
}
{
right.
destruct H2 as [E|[E|E]].
{
subst ty'.
simpl typesize.
omega.
}
{
subst ty'.
simpl typesize.
omega.
}
{
subst ty'.
simpl typesize.
omega.
}
}
}
}
}
{
destruct H.
{
left.
apply OrderedSlot.lt_not_eq.
auto.
}
{
destruct H.
right.
destruct H0.
{
left.
omega.
}
{
destruct H0.
exfalso.
destruct ty'.
{
compute in H1.
congruence.
}
{
compute in H1.
congruence.
}
{
compute in H1.
congruence.
}
{
compute in H1.
congruence.
}
{
compute in H1.
congruence.
}
{
compute in H1.
congruence.
}
}
}
}
}
}
}
Qed.
Proof.
Admitted.
Proof.
intros.
destruct l as [mr | sl ofs ty].
{
destruct l' as [mr' | sl' ofs' ty'].
{
simpl in *.
auto.
unfold Plt.
unfold Pos.lt.
destruct (Pos.compare (IndexedMreg.index mr) (IndexedMreg.index mr')) eqn:C.
{
elim H.
apply IndexedMreg.index_inj.
apply Pos.compare_eq_iff.
auto.
}
{
auto.
}
{
rewrite Pos.compare_antisym.
rewrite C.
auto.
}
}
{
simpl in *.
auto.
}
}
{
destruct l' as [mr' | sl' ofs' ty'].
{
simpl in *.
auto.
}
{
simpl in *.
auto.
destruct (OrderedSlot.compare sl sl').
{
auto.
}
{
auto.
destruct H.
{
contradiction.
}
{
destruct H.
{
right.
right.
split.
{
auto.
}
{
auto.
left.
omega.
}
}
{
left.
right.
split.
{
auto.
}
{
auto.
assert (EITHER: typesize ty' = 1 /\ OrderedTyp.lt ty' Tany64 \/ typesize ty' = 2).
{
destruct ty'.
{
compute.
auto.
}
{
compute.
auto.
}
{
compute.
auto.
}
{
compute.
auto.
}
{
compute.
auto.
}
{
compute.
auto.
}
}
{
destruct (zlt ofs' (ofs - 1)).
{
left.
auto.
}
{
destruct EITHER as [[P Q] | P].
{
right.
split.
{
auto.
omega.
}
{
auto.
}
}
{
left.
omega.
}
}
}
}
}
}
}
{
auto.
}
}
}
Qed.
End OrderedLoc.