Library compcert.cfrontend.Cop


Arithmetic and logical operators for the Compcert C and Clight languages

Require Import Coqlib.
Require Import AST.
Require Import Integers.
Require Import Floats.
Require Import Values.
Require Import Memory.
Require Import Ctypes.
Require Archi.

Syntax of operators.

Type classification and semantics of operators.

Most C operators are overloaded (they apply to arguments of various types) and their semantics depend on the types of their arguments. The following classify_× functions take as arguments the types of the arguments of an operation. They return enough information to resolve overloading for this operator applications, such as ``both arguments are floats'', or ``the first is a pointer and the second is an integer''. This classification is used in the compiler (module Cshmgen) to resolve overloading statically.
The sem_× functions below compute the result of an operator application. Since operators are overloaded, the result depends both on the static types of the arguments and on their run-time values. The corresponding classify_× function is first called on the types of the arguments to resolve static overloading. It is then followed by a case analysis on the values of the arguments.

Casts and truth values


Inductive classify_cast_cases : Type :=
  | cast_case_pointer
  | cast_case_i2i (sz2:intsize) (si2:signedness)
  | cast_case_f2f
  | cast_case_s2s
  | cast_case_f2s
  | cast_case_s2f
  | cast_case_i2f (si1: signedness)
  | cast_case_i2s (si1: signedness)
  | cast_case_f2i (sz2:intsize) (si2:signedness)
  | cast_case_s2i (sz2:intsize) (si2:signedness)
  | cast_case_l2l
  | cast_case_i2l (si1: signedness)
  | cast_case_l2i (sz2: intsize) (si2: signedness)
  | cast_case_l2f (si1: signedness)
  | cast_case_l2s (si1: signedness)
  | cast_case_f2l (si2:signedness)
  | cast_case_s2l (si2:signedness)
  | cast_case_i2bool
  | cast_case_l2bool
  | cast_case_f2bool
  | cast_case_s2bool
  | cast_case_struct (id1 id2: ident)
  | cast_case_union (id1 id2: ident)
  | cast_case_void
  | cast_case_default.

Definition classify_cast (tfrom tto: type) : classify_cast_cases :=
  match tto, tfrom with
  
  | Tvoid, _cast_case_void
  
  | Tint IBool _ _, Tint _ _ _cast_case_i2bool
  | Tint IBool _ _, Tlong _ _cast_case_l2bool
  | Tint IBool _ _, Tfloat F64 _cast_case_f2bool
  | Tint IBool _ _, Tfloat F32 _cast_case_s2bool
  | Tint IBool _ _, (Tpointer _ _ | Tarray _ _ _ | Tfunction _ _ _) ⇒
      if Archi.ptr64 then cast_case_l2bool else cast_case_i2bool
  
  | Tint sz2 si2 _, Tint _ _ _
      if Archi.ptr64 then cast_case_i2i sz2 si2
      else if intsize_eq sz2 I32 then cast_case_pointer
      else cast_case_i2i sz2 si2
  | Tint sz2 si2 _, Tlong _ _cast_case_l2i sz2 si2
  | Tint sz2 si2 _, Tfloat F64 _cast_case_f2i sz2 si2
  | Tint sz2 si2 _, Tfloat F32 _cast_case_s2i sz2 si2
  | Tint sz2 si2 _, (Tpointer _ _ | Tarray _ _ _ | Tfunction _ _ _) ⇒
      if Archi.ptr64 then cast_case_l2i sz2 si2
      else if intsize_eq sz2 I32 then cast_case_pointer
      else cast_case_i2i sz2 si2
  
  | Tlong _ _, Tlong _ _
      if Archi.ptr64 then cast_case_pointer else cast_case_l2l
  | Tlong _ _, Tint sz1 si1 _cast_case_i2l si1
  | Tlong si2 _, Tfloat F64 _cast_case_f2l si2
  | Tlong si2 _, Tfloat F32 _cast_case_s2l si2
  | Tlong si2 _, (Tpointer _ _ | Tarray _ _ _ | Tfunction _ _ _) ⇒
      if Archi.ptr64 then cast_case_pointer else cast_case_i2l si2
  
  | Tfloat F64 _, Tint sz1 si1 _cast_case_i2f si1
  | Tfloat F32 _, Tint sz1 si1 _cast_case_i2s si1
  | Tfloat F64 _, Tlong si1 _cast_case_l2f si1
  | Tfloat F32 _, Tlong si1 _cast_case_l2s si1
  | Tfloat F64 _, Tfloat F64 _cast_case_f2f
  | Tfloat F32 _, Tfloat F32 _cast_case_s2s
  | Tfloat F64 _, Tfloat F32 _cast_case_s2f
  | Tfloat F32 _, Tfloat F64 _cast_case_f2s
  
  | Tpointer _ _, Tint _ _ _
      if Archi.ptr64 then cast_case_i2l Unsigned else cast_case_pointer
  | Tpointer _ _, Tlong _ _
      if Archi.ptr64 then cast_case_pointer else cast_case_l2i I32 Unsigned
  | Tpointer _ _, (Tpointer _ _ | Tarray _ _ _ | Tfunction _ _ _) ⇒ cast_case_pointer
  
  | Tstruct id2 _, Tstruct id1 _cast_case_struct id1 id2
  | Tunion id2 _, Tunion id1 _cast_case_union id1 id2
  
  | _, _cast_case_default
  end.

Semantics of casts. sem_cast v1 t1 t2 m = Some v2 if value v1, viewed with static type t1, can be converted to type t2, resulting in value v2.

Definition cast_int_int (sz: intsize) (sg: signedness) (i: int) : int :=
  match sz, sg with
  | I8, SignedInt.sign_ext 8 i
  | I8, UnsignedInt.zero_ext 8 i
  | I16, SignedInt.sign_ext 16 i
  | I16, UnsignedInt.zero_ext 16 i
  | I32, _i
  | IBool, _if Int.eq i Int.zero then Int.zero else Int.one
  end.

Definition cast_int_float (si: signedness) (i: int) : float :=
  match si with
  | SignedFloat.of_int i
  | UnsignedFloat.of_intu i
  end.

Definition cast_float_int (si : signedness) (f: float) : option int :=
  match si with
  | SignedFloat.to_int f
  | UnsignedFloat.to_intu f
  end.

Definition cast_int_single (si: signedness) (i: int) : float32 :=
  match si with
  | SignedFloat32.of_int i
  | UnsignedFloat32.of_intu i
  end.

Definition cast_single_int (si : signedness) (f: float32) : option int :=
  match si with
  | SignedFloat32.to_int f
  | UnsignedFloat32.to_intu f
  end.

Definition cast_int_long (si: signedness) (i: int) : int64 :=
  match si with
  | SignedInt64.repr (Int.signed i)
  | UnsignedInt64.repr (Int.unsigned i)
  end.

Definition cast_long_float (si: signedness) (i: int64) : float :=
  match si with
  | SignedFloat.of_long i
  | UnsignedFloat.of_longu i
  end.

Definition cast_long_single (si: signedness) (i: int64) : float32 :=
  match si with
  | SignedFloat32.of_long i
  | UnsignedFloat32.of_longu i
  end.

Definition cast_float_long (si : signedness) (f: float) : option int64 :=
  match si with
  | SignedFloat.to_long f
  | UnsignedFloat.to_longu f
  end.

Definition cast_single_long (si : signedness) (f: float32) : option int64 :=
  match si with
  | SignedFloat32.to_long f
  | UnsignedFloat32.to_longu f
  end.


Class SemCast {T: Type} (valid_pointer: T block Z bool): Prop :=
  {
    weak_valid_pointer m b o :=
      valid_pointer m b o || valid_pointer m b (o - 1)
  }.

Section WITHSEMCAST.
Context `{sem_cast_prf: SemCast}.

Definition sem_cast (v: val) (t1 t2: type) (m: T): option val :=
  match classify_cast t1 t2 with
  | cast_case_pointer
      match v with
      | Vptr _ _Some v
      | Vint _if Archi.ptr64 then None else Some v
      | Vlong _if Archi.ptr64 then Some v else None
      | _None
      end
  | cast_case_i2i sz2 si2
      match v with
      | Vint iSome (Vint (cast_int_int sz2 si2 i))
      | _None
      end
  | cast_case_f2f
      match v with
      | Vfloat fSome (Vfloat f)
      | _None
      end
  | cast_case_s2s
      match v with
      | Vsingle fSome (Vsingle f)
      | _None
      end
  | cast_case_s2f
      match v with
      | Vsingle fSome (Vfloat (Float.of_single f))
      | _None
      end
  | cast_case_f2s
      match v with
      | Vfloat fSome (Vsingle (Float.to_single f))
      | _None
      end
  | cast_case_i2f si1
      match v with
      | Vint iSome (Vfloat (cast_int_float si1 i))
      | _None
      end
  | cast_case_i2s si1
      match v with
      | Vint iSome (Vsingle (cast_int_single si1 i))
      | _None
      end
  | cast_case_f2i sz2 si2
      match v with
      | Vfloat f
          match cast_float_int si2 f with
          | Some iSome (Vint (cast_int_int sz2 si2 i))
          | NoneNone
          end
      | _None
      end
  | cast_case_s2i sz2 si2
      match v with
      | Vsingle f
          match cast_single_int si2 f with
          | Some iSome (Vint (cast_int_int sz2 si2 i))
          | NoneNone
          end
      | _None
      end
  | cast_case_i2bool
      match v with
      | Vint n
          Some(Vint(if Int.eq n Int.zero then Int.zero else Int.one))
      | Vptr b ofs
          if Archi.ptr64 then None else
          if weak_valid_pointer m b (Ptrofs.unsigned ofs) then Some Vone else None
      | _None
      end
  | cast_case_l2bool
      match v with
      | Vlong n
          Some(Vint(if Int64.eq n Int64.zero then Int.zero else Int.one))
      | Vptr b ofs
          if negb Archi.ptr64 then None else
            if weak_valid_pointer m b (Ptrofs.unsigned ofs) then Some Vone else None
      | _None
      end
  | cast_case_f2bool
      match v with
      | Vfloat f
          Some(Vint(if Float.cmp Ceq f Float.zero then Int.zero else Int.one))
      | _None
      end
  | cast_case_s2bool
      match v with
      | Vsingle f
          Some(Vint(if Float32.cmp Ceq f Float32.zero then Int.zero else Int.one))
      | _None
      end
  | cast_case_l2l
      match v with
      | Vlong nSome (Vlong n)
      | _None
      end
  | cast_case_i2l si
      match v with
      | Vint nSome(Vlong (cast_int_long si n))
      | _None
      end
  | cast_case_l2i sz si
      match v with
      | Vlong nSome(Vint (cast_int_int sz si (Int.repr (Int64.unsigned n))))
      | _None
      end
  | cast_case_l2f si1
      match v with
      | Vlong iSome (Vfloat (cast_long_float si1 i))
      | _None
      end
  | cast_case_l2s si1
      match v with
      | Vlong iSome (Vsingle (cast_long_single si1 i))
      | _None
      end
  | cast_case_f2l si2
      match v with
      | Vfloat f
          match cast_float_long si2 f with
          | Some iSome (Vlong i)
          | NoneNone
          end
      | _None
      end
  | cast_case_s2l si2
      match v with
      | Vsingle f
          match cast_single_long si2 f with
          | Some iSome (Vlong i)
          | NoneNone
          end
      | _None
      end
  | cast_case_struct id1 id2
      match v with
      | Vptr b ofs
          if ident_eq id1 id2 then Some v else None
      | _None
      end
  | cast_case_union id1 id2
      match v with
      | Vptr b ofs
          if ident_eq id1 id2 then Some v else None
      | _None
      end
  | cast_case_void
      Some v
  | cast_case_default
      None
  end.

The following describes types that can be interpreted as a boolean: integers, floats, pointers. It is used for the semantics of the ! and ? operators, as well as the if, while, and for statements.

Inductive classify_bool_cases : Type :=
  | bool_case_i
  | bool_case_l
  | bool_case_f
  | bool_case_s
  | bool_default.

Definition classify_bool (ty: type) : classify_bool_cases :=
  match typeconv ty with
  | Tint _ _ _bool_case_i
  | Tpointer _ _if Archi.ptr64 then bool_case_l else bool_case_i
  | Tfloat F64 _bool_case_f
  | Tfloat F32 _bool_case_s
  | Tlong _ _bool_case_l
  | _bool_default
  end.

Interpretation of values as truth values. Non-zero integers, non-zero floats and non-null pointers are considered as true. The integer zero (which also represents the null pointer) and the float 0.0 are false.

Definition bool_val (v: val) (t: type) (m: T) : option bool :=
  match classify_bool t with
  | bool_case_i
      match v with
      | Vint nSome (negb (Int.eq n Int.zero))
      | Vptr b ofs
          if Archi.ptr64 then None else
          if weak_valid_pointer m b (Ptrofs.unsigned ofs) then Some true else None
      | _None
      end
  | bool_case_l
      match v with
      | Vlong nSome (negb (Int64.eq n Int64.zero))
      | Vptr b ofs
          if negb Archi.ptr64 then None else
          if weak_valid_pointer m b (Ptrofs.unsigned ofs) then Some true else None
      | _None
      end
  | bool_case_f
      match v with
      | Vfloat fSome (negb (Float.cmp Ceq f Float.zero))
      | _None
      end
  | bool_case_s
      match v with
      | Vsingle fSome (negb (Float32.cmp Ceq f Float32.zero))
      | _None
      end
  | bool_defaultNone
  end.

Unary operators

Boolean negation


Definition sem_notbool (v: val) (ty: type) (m: T): option val :=
  option_map (fun bVal.of_bool (negb b)) (bool_val v ty m).

Opposite and absolute value


Inductive classify_neg_cases : Type :=
  | neg_case_i(s: signedness)
  | neg_case_f
  | neg_case_s
  | neg_case_l(s: signedness)
  | neg_default.

Definition classify_neg (ty: type) : classify_neg_cases :=
  match ty with
  | Tint I32 Unsigned _neg_case_i Unsigned
  | Tint _ _ _neg_case_i Signed
  | Tfloat F64 _neg_case_f
  | Tfloat F32 _neg_case_s
  | Tlong si _neg_case_l si
  | _neg_default
  end.

Definition sem_neg (v: val) (ty: type) : option val :=
  match classify_neg ty with
  | neg_case_i sg
      match v with
      | Vint nSome (Vint (Int.neg n))
      | _None
      end
  | neg_case_f
      match v with
      | Vfloat fSome (Vfloat (Float.neg f))
      | _None
      end
  | neg_case_s
      match v with
      | Vsingle fSome (Vsingle (Float32.neg f))
      | _None
      end
  | neg_case_l sg
      match v with
      | Vlong nSome (Vlong (Int64.neg n))
      | _None
      end
  | neg_defaultNone
  end.

Definition sem_absfloat (v: val) (ty: type) : option val :=
  match classify_neg ty with
  | neg_case_i sg
      match v with
      | Vint nSome (Vfloat (Float.abs (cast_int_float sg n)))
      | _None
      end
  | neg_case_f
      match v with
      | Vfloat fSome (Vfloat (Float.abs f))
      | _None
      end
  | neg_case_s
      match v with
      | Vsingle fSome (Vfloat (Float.abs (Float.of_single f)))
      | _None
      end
  | neg_case_l sg
      match v with
      | Vlong nSome (Vfloat (Float.abs (cast_long_float sg n)))
      | _None
      end
  | neg_defaultNone
  end.

Bitwise complement


Inductive classify_notint_cases : Type :=
  | notint_case_i(s: signedness)
  | notint_case_l(s: signedness)
  | notint_default.

Definition classify_notint (ty: type) : classify_notint_cases :=
  match ty with
  | Tint I32 Unsigned _notint_case_i Unsigned
  | Tint _ _ _notint_case_i Signed
  | Tlong si _notint_case_l si
  | _notint_default
  end.

Definition sem_notint (v: val) (ty: type): option val :=
  match classify_notint ty with
  | notint_case_i sg
      match v with
      | Vint nSome (Vint (Int.not n))
      | _None
      end
  | notint_case_l sg
      match v with
      | Vlong nSome (Vlong (Int64.not n))
      | _None
      end
  | notint_defaultNone
  end.

Binary operators

For binary operations, the "usual binary conversions" consist in
  • determining the type at which the operation is to be performed (a form of least upper bound of the types of the two arguments);
  • casting the two arguments to this common type;
  • performing the operation at that type.

Inductive binarith_cases: Type :=
  | bin_case_i (s: signedness)
  | bin_case_l (s: signedness)
  | bin_case_f
  | bin_case_s
  | bin_default.
Definition classify_binarith (ty1: type) (ty2: type) : binarith_cases :=
  match ty1, ty2 with
  | Tint I32 Unsigned _, Tint _ _ _bin_case_i Unsigned
  | Tint _ _ _, Tint I32 Unsigned _bin_case_i Unsigned
  | Tint _ _ _, Tint _ _ _bin_case_i Signed
  | Tlong Signed _, Tlong Signed _bin_case_l Signed
  | Tlong _ _, Tlong _ _bin_case_l Unsigned
  | Tlong sg _, Tint _ _ _bin_case_l sg
  | Tint _ _ _, Tlong sg _bin_case_l sg
  | Tfloat F32 _, Tfloat F32 _bin_case_s
  | Tfloat _ _, Tfloat _ _bin_case_f
  | Tfloat F64 _, (Tint _ _ _ | Tlong _ _) ⇒ bin_case_f
  | (Tint _ _ _ | Tlong _ _), Tfloat F64 _bin_case_f
  | Tfloat F32 _, (Tint _ _ _ | Tlong _ _) ⇒ bin_case_s
  | (Tint _ _ _ | Tlong _ _), Tfloat F32 _bin_case_s
  | _, _bin_default
  end.

The static type of the result. Both arguments are converted to this type before the actual computation.

Definition binarith_type (c: binarith_cases) : type :=
  match c with
  | bin_case_i sgTint I32 sg noattr
  | bin_case_l sgTlong sg noattr
  | bin_case_fTfloat F64 noattr
  | bin_case_sTfloat F32 noattr
  | bin_defaultTvoid
  end.

Definition sem_binarith
    (sem_int: signedness int int option val)
    (sem_long: signedness int64 int64 option val)
    (sem_float: float float option val)
    (sem_single: float32 float32 option val)
    (v1: val) (t1: type) (v2: val) (t2: type) (m: T): option val :=
  let c := classify_binarith t1 t2 in
  let t := binarith_type c in
  match sem_cast v1 t1 t m with
  | NoneNone
  | Some v1'
  match sem_cast v2 t2 t m with
  | NoneNone
  | Some v2'
  match c with
  | bin_case_i sg
      match v1', v2' with
      | Vint n1, Vint n2sem_int sg n1 n2
      | _, _None
      end
  | bin_case_f
      match v1', v2' with
      | Vfloat n1, Vfloat n2sem_float n1 n2
      | _, _None
      end
  | bin_case_s
      match v1', v2' with
      | Vsingle n1, Vsingle n2sem_single n1 n2
      | _, _None
      end
  | bin_case_l sg
      match v1', v2' with
      | Vlong n1, Vlong n2sem_long sg n1 n2
      | _, _None
      end
  | bin_defaultNone
  end end end.

Addition


Inductive classify_add_cases : Type :=
  | add_case_pi (ty: type) (si: signedness)
  | add_case_pl (ty: type)
  | add_case_ip (si: signedness) (ty: type)
  | add_case_lp (ty: type)
  | add_default.
Definition classify_add (ty1: type) (ty2: type) :=
  match typeconv ty1, typeconv ty2 with
  | Tpointer ty _, Tint _ si _add_case_pi ty si
  | Tpointer ty _, Tlong _ _add_case_pl ty
  | Tint _ si _, Tpointer ty _add_case_ip si ty
  | Tlong _ _, Tpointer ty _add_case_lp ty
  | _, _add_default
  end.

Definition ptrofs_of_int (si: signedness) (n: int) : ptrofs :=
  match si with
  | SignedPtrofs.of_ints n
  | UnsignedPtrofs.of_intu n
  end.

Definition sem_add_ptr_int (cenv: composite_env) (ty: type) (si: signedness) (v1 v2: val): option val :=
  match v1, v2 with
  | Vptr b1 ofs1, Vint n2
      let n2 := ptrofs_of_int si n2 in
      Some (Vptr b1 (Ptrofs.add ofs1 (Ptrofs.mul (Ptrofs.repr (sizeof cenv ty)) n2)))
  | Vint n1, Vint n2
      if Archi.ptr64 then None else Some (Vint (Int.add n1 (Int.mul (Int.repr (sizeof cenv ty)) n2)))
  | Vlong n1, Vint n2
      let n2 := cast_int_long si n2 in
      if Archi.ptr64 then Some (Vlong (Int64.add n1 (Int64.mul (Int64.repr (sizeof cenv ty)) n2))) else None
  | _, _None
  end.

Definition sem_add_ptr_long (cenv: composite_env) (ty: type) (v1 v2: val): option val :=
  match v1, v2 with
  | Vptr b1 ofs1, Vlong n2
      let n2 := Ptrofs.of_int64 n2 in
      Some (Vptr b1 (Ptrofs.add ofs1 (Ptrofs.mul (Ptrofs.repr (sizeof cenv ty)) n2)))
  | Vint n1, Vlong n2
      let n2 := Int.repr (Int64.unsigned n2) in
      if Archi.ptr64 then None else Some (Vint (Int.add n1 (Int.mul (Int.repr (sizeof cenv ty)) n2)))
  | Vlong n1, Vlong n2
      if Archi.ptr64 then Some (Vlong (Int64.add n1 (Int64.mul (Int64.repr (sizeof cenv ty)) n2))) else None
  | _, _None
  end.

Definition sem_add (cenv: composite_env) (v1:val) (t1:type) (v2: val) (t2:type) (m: T): option val :=
  match classify_add t1 t2 with
  | add_case_pi ty si
      sem_add_ptr_int cenv ty si v1 v2
  | add_case_pl ty
      sem_add_ptr_long cenv ty v1 v2
  | add_case_ip si ty
      sem_add_ptr_int cenv ty si v2 v1
  | add_case_lp ty
      sem_add_ptr_long cenv ty v2 v1
  | add_default
      sem_binarith
        (fun sg n1 n2Some(Vint(Int.add n1 n2)))
        (fun sg n1 n2Some(Vlong(Int64.add n1 n2)))
        (fun n1 n2Some(Vfloat(Float.add n1 n2)))
        (fun n1 n2Some(Vsingle(Float32.add n1 n2)))
        v1 t1 v2 t2 m
  end.

Subtraction


Inductive classify_sub_cases : Type :=
  | sub_case_pi (ty: type) (si: signedness)
  | sub_case_pp (ty: type)
  | sub_case_pl (ty: type)
  | sub_default.
Definition classify_sub (ty1: type) (ty2: type) :=
  match typeconv ty1, typeconv ty2 with
  | Tpointer ty _, Tint _ si _sub_case_pi ty si
  | Tpointer ty _ , Tpointer _ _sub_case_pp ty
  | Tpointer ty _, Tlong _ _sub_case_pl ty
  | _, _sub_default
  end.

Definition sem_sub (cenv: composite_env) (v1:val) (t1:type) (v2: val) (t2:type) (m:T): option val :=
  match classify_sub t1 t2 with
  | sub_case_pi ty si
      match v1, v2 with
      | Vptr b1 ofs1, Vint n2
          let n2 := ptrofs_of_int si n2 in
          Some (Vptr b1 (Ptrofs.sub ofs1 (Ptrofs.mul (Ptrofs.repr (sizeof cenv ty)) n2)))
      | Vint n1, Vint n2
          if Archi.ptr64 then None else Some (Vint (Int.sub n1 (Int.mul (Int.repr (sizeof cenv ty)) n2)))
      | Vlong n1, Vint n2
          let n2 := cast_int_long si n2 in
          if Archi.ptr64 then Some (Vlong (Int64.sub n1 (Int64.mul (Int64.repr (sizeof cenv ty)) n2))) else None
      | _, _None
      end
  | sub_case_pl ty
      match v1, v2 with
      | Vptr b1 ofs1, Vlong n2
          let n2 := Ptrofs.of_int64 n2 in
          Some (Vptr b1 (Ptrofs.sub ofs1 (Ptrofs.mul (Ptrofs.repr (sizeof cenv ty)) n2)))
      | Vint n1, Vlong n2
          let n2 := Int.repr (Int64.unsigned n2) in
          if Archi.ptr64 then None else Some (Vint (Int.sub n1 (Int.mul (Int.repr (sizeof cenv ty)) n2)))
      | Vlong n1, Vlong n2
          if Archi.ptr64 then Some (Vlong (Int64.sub n1 (Int64.mul (Int64.repr (sizeof cenv ty)) n2))) else None
      | _, _None
      end
  | sub_case_pp ty
      match v1,v2 with
      | Vptr b1 ofs1, Vptr b2 ofs2
          if eq_block b1 b2 then
            let sz := sizeof cenv ty in
            if zlt 0 sz && zle sz Ptrofs.max_signed
            then Some (Vptrofs (Ptrofs.divs (Ptrofs.sub ofs1 ofs2) (Ptrofs.repr sz)))
            else None
          else None
      | _, _None
      end
  | sub_default
      sem_binarith
        (fun sg n1 n2Some(Vint(Int.sub n1 n2)))
        (fun sg n1 n2Some(Vlong(Int64.sub n1 n2)))
        (fun n1 n2Some(Vfloat(Float.sub n1 n2)))
        (fun n1 n2Some(Vsingle(Float32.sub n1 n2)))
        v1 t1 v2 t2 m
  end.

Multiplication, division, modulus


Definition sem_mul (v1:val) (t1:type) (v2: val) (t2:type) (m:T) : option val :=
  sem_binarith
    (fun sg n1 n2Some(Vint(Int.mul n1 n2)))
    (fun sg n1 n2Some(Vlong(Int64.mul n1 n2)))
    (fun n1 n2Some(Vfloat(Float.mul n1 n2)))
    (fun n1 n2Some(Vsingle(Float32.mul n1 n2)))
    v1 t1 v2 t2 m.

Definition sem_div (v1:val) (t1:type) (v2: val) (t2:type) (m:T) : option val :=
  sem_binarith
    (fun sg n1 n2
      match sg with
      | Signed
          if Int.eq n2 Int.zero
          || Int.eq n1 (Int.repr Int.min_signed) && Int.eq n2 Int.mone
          then None else Some(Vint(Int.divs n1 n2))
      | Unsigned
          if Int.eq n2 Int.zero
          then None else Some(Vint(Int.divu n1 n2))
      end)
    (fun sg n1 n2
      match sg with
      | Signed
          if Int64.eq n2 Int64.zero
          || Int64.eq n1 (Int64.repr Int64.min_signed) && Int64.eq n2 Int64.mone
          then None else Some(Vlong(Int64.divs n1 n2))
      | Unsigned
          if Int64.eq n2 Int64.zero
          then None else Some(Vlong(Int64.divu n1 n2))
      end)
    (fun n1 n2Some(Vfloat(Float.div n1 n2)))
    (fun n1 n2Some(Vsingle(Float32.div n1 n2)))
    v1 t1 v2 t2 m.

Definition sem_mod (v1:val) (t1:type) (v2: val) (t2:type) (m:T) : option val :=
  sem_binarith
    (fun sg n1 n2
      match sg with
      | Signed
          if Int.eq n2 Int.zero
          || Int.eq n1 (Int.repr Int.min_signed) && Int.eq n2 Int.mone
          then None else Some(Vint(Int.mods n1 n2))
      | Unsigned
          if Int.eq n2 Int.zero
          then None else Some(Vint(Int.modu n1 n2))
      end)
    (fun sg n1 n2
      match sg with
      | Signed
          if Int64.eq n2 Int64.zero
          || Int64.eq n1 (Int64.repr Int64.min_signed) && Int64.eq n2 Int64.mone
          then None else Some(Vlong(Int64.mods n1 n2))
      | Unsigned
          if Int64.eq n2 Int64.zero
          then None else Some(Vlong(Int64.modu n1 n2))
      end)
    (fun n1 n2None)
    (fun n1 n2None)
    v1 t1 v2 t2 m.

Definition sem_and (v1:val) (t1:type) (v2: val) (t2:type) (m:T) : option val :=
  sem_binarith
    (fun sg n1 n2Some(Vint(Int.and n1 n2)))
    (fun sg n1 n2Some(Vlong(Int64.and n1 n2)))
    (fun n1 n2None)
    (fun n1 n2None)
    v1 t1 v2 t2 m.

Definition sem_or (v1:val) (t1:type) (v2: val) (t2:type) (m:T) : option val :=
  sem_binarith
    (fun sg n1 n2Some(Vint(Int.or n1 n2)))
    (fun sg n1 n2Some(Vlong(Int64.or n1 n2)))
    (fun n1 n2None)
    (fun n1 n2None)
    v1 t1 v2 t2 m.

Definition sem_xor (v1:val) (t1:type) (v2: val) (t2:type) (m:T) : option val :=
  sem_binarith
    (fun sg n1 n2Some(Vint(Int.xor n1 n2)))
    (fun sg n1 n2Some(Vlong(Int64.xor n1 n2)))
    (fun n1 n2None)
    (fun n1 n2None)
    v1 t1 v2 t2 m.

Shifts

Shifts do not perform the usual binary conversions. Instead, each argument is converted independently, and the signedness of the result is always that of the first argument.

Inductive classify_shift_cases : Type:=
  | shift_case_ii(s: signedness)
  | shift_case_ll(s: signedness)
  | shift_case_il(s: signedness)
  | shift_case_li(s: signedness)
  | shift_default.

Definition classify_shift (ty1: type) (ty2: type) :=
  match typeconv ty1, typeconv ty2 with
  | Tint I32 Unsigned _, Tint _ _ _shift_case_ii Unsigned
  | Tint _ _ _, Tint _ _ _shift_case_ii Signed
  | Tint I32 Unsigned _, Tlong _ _shift_case_il Unsigned
  | Tint _ _ _, Tlong _ _shift_case_il Signed
  | Tlong s _, Tint _ _ _shift_case_li s
  | Tlong s _, Tlong _ _shift_case_ll s
  | _,_shift_default
  end.

Definition sem_shift
    (sem_int: signedness int int int)
    (sem_long: signedness int64 int64 int64)
    (v1: val) (t1: type) (v2: val) (t2: type) : option val :=
  match classify_shift t1 t2 with
  | shift_case_ii sg
      match v1, v2 with
      | Vint n1, Vint n2
          if Int.ltu n2 Int.iwordsize
          then Some(Vint(sem_int sg n1 n2)) else None
      | _, _None
      end
  | shift_case_il sg
      match v1, v2 with
      | Vint n1, Vlong n2
          if Int64.ltu n2 (Int64.repr 32)
          then Some(Vint(sem_int sg n1 (Int64.loword n2))) else None
      | _, _None
      end
  | shift_case_li sg
      match v1, v2 with
      | Vlong n1, Vint n2
          if Int.ltu n2 Int64.iwordsize'
          then Some(Vlong(sem_long sg n1 (Int64.repr (Int.unsigned n2)))) else None
      | _, _None
      end
  | shift_case_ll sg
      match v1, v2 with
      | Vlong n1, Vlong n2
          if Int64.ltu n2 Int64.iwordsize
          then Some(Vlong(sem_long sg n1 n2)) else None
      | _, _None
      end
  | shift_defaultNone
  end.

Definition sem_shl (v1:val) (t1:type) (v2: val) (t2:type) : option val :=
  sem_shift
    (fun sg n1 n2Int.shl n1 n2)
    (fun sg n1 n2Int64.shl n1 n2)
    v1 t1 v2 t2.

Definition sem_shr (v1:val) (t1:type) (v2: val) (t2:type) : option val :=
  sem_shift
    (fun sg n1 n2match sg with SignedInt.shr n1 n2 | UnsignedInt.shru n1 n2 end)
    (fun sg n1 n2match sg with SignedInt64.shr n1 n2 | UnsignedInt64.shru n1 n2 end)
    v1 t1 v2 t2.

Comparisons


Inductive classify_cmp_cases : Type :=
  | cmp_case_pp
  | cmp_case_pi (si: signedness)
  | cmp_case_ip (si: signedness)
  | cmp_case_pl
  | cmp_case_lp
  | cmp_default.
Definition classify_cmp (ty1: type) (ty2: type) :=
  match typeconv ty1, typeconv ty2 with
  | Tpointer _ _ , Tpointer _ _cmp_case_pp
  | Tpointer _ _ , Tint _ si _cmp_case_pi si
  | Tint _ si _, Tpointer _ _cmp_case_ip si
  | Tpointer _ _ , Tlong _ _cmp_case_pl
  | Tlong _ _ , Tpointer _ _cmp_case_lp
  | _, _cmp_default
  end.

Definition cmp_ptr (m: T) (c: comparison) (v1 v2: val): option val :=
  option_map Val.of_bool
   (if Archi.ptr64
    then Val.cmplu_bool (valid_pointer m) c v1 v2
    else Val.cmpu_bool (valid_pointer m) c v1 v2).

Definition sem_cmp (c:comparison)
                  (v1: val) (t1: type) (v2: val) (t2: type)
                  (m: T): option val :=
  match classify_cmp t1 t2 with
  | cmp_case_pp
      cmp_ptr m c v1 v2
  | cmp_case_pi si
      match v2 with
      | Vint n2
          let v2' := Vptrofs (ptrofs_of_int si n2) in
          cmp_ptr m c v1 v2'
      | Vptr b ofs
          if Archi.ptr64 then None else cmp_ptr m c v1 v2
      | _
          None
      end
  | cmp_case_ip si
      match v1 with
      | Vint n1
          let v1' := Vptrofs (ptrofs_of_int si n1) in
          cmp_ptr m c v1' v2
      | Vptr b ofs
          if Archi.ptr64 then None else cmp_ptr m c v1 v2
      | _
          None
      end
  | cmp_case_pl
      match v2 with
      | Vlong n2
          let v2' := Vptrofs (Ptrofs.of_int64 n2) in
          cmp_ptr m c v1 v2'
      | Vptr b ofs
          if Archi.ptr64 then cmp_ptr m c v1 v2 else None
      | _
          None
      end
  | cmp_case_lp
      match v1 with
      | Vlong n1
          let v1' := Vptrofs (Ptrofs.of_int64 n1) in
          cmp_ptr m c v1' v2
      | Vptr b ofs
          if Archi.ptr64 then cmp_ptr m c v1 v2 else None
      | _
          None
      end
  | cmp_default
      sem_binarith
        (fun sg n1 n2
            Some(Val.of_bool(match sg with SignedInt.cmp c n1 n2 | UnsignedInt.cmpu c n1 n2 end)))
        (fun sg n1 n2
            Some(Val.of_bool(match sg with SignedInt64.cmp c n1 n2 | UnsignedInt64.cmpu c n1 n2 end)))
        (fun n1 n2
            Some(Val.of_bool(Float.cmp c n1 n2)))
        (fun n1 n2
            Some(Val.of_bool(Float32.cmp c n1 n2)))
        v1 t1 v2 t2 m
  end.

Function applications


Inductive classify_fun_cases : Type :=
  | fun_case_f (targs: typelist) (tres: type) (cc: calling_convention)
  | fun_default.

Definition classify_fun (ty: type) :=
  match ty with
  | Tfunction args res ccfun_case_f args res cc
  | Tpointer (Tfunction args res cc) _fun_case_f args res cc
  | _fun_default
  end.

Argument of a switch statement


Inductive classify_switch_cases : Type :=
  | switch_case_i
  | switch_case_l
  | switch_default.

Definition classify_switch (ty: type) :=
  match ty with
  | Tint _ _ _switch_case_i
  | Tlong _ _switch_case_l
  | _switch_default
  end.

Definition sem_switch_arg (v: val) (ty: type): option Z :=
  match classify_switch ty with
  | switch_case_i
      match v with Vint nSome(Int.unsigned n) | _None end
  | switch_case_l
      match v with Vlong nSome(Int64.unsigned n) | _None end
  | switch_default
      None
  end.

Combined semantics of unary and binary operators


Definition sem_unary_operation
            (op: unary_operation) (v: val) (ty: type) (m: T): option val :=
  match op with
  | Onotboolsem_notbool v ty m
  | Onotintsem_notint v ty
  | Onegsem_neg v ty
  | Oabsfloatsem_absfloat v ty
  end.

Definition sem_binary_operation
    (cenv: composite_env)
    (op: binary_operation)
    (v1: val) (t1: type) (v2: val) (t2:type)
    (m: T): option val :=
  match op with
  | Oaddsem_add cenv v1 t1 v2 t2 m
  | Osubsem_sub cenv v1 t1 v2 t2 m
  | Omulsem_mul v1 t1 v2 t2 m
  | Omodsem_mod v1 t1 v2 t2 m
  | Odivsem_div v1 t1 v2 t2 m
  | Oandsem_and v1 t1 v2 t2 m
  | Oorsem_or v1 t1 v2 t2 m
  | Oxorsem_xor v1 t1 v2 t2 m
  | Oshlsem_shl v1 t1 v2 t2
  | Oshrsem_shr v1 t1 v2 t2
  | Oeqsem_cmp Ceq v1 t1 v2 t2 m
  | Onesem_cmp Cne v1 t1 v2 t2 m
  | Oltsem_cmp Clt v1 t1 v2 t2 m
  | Ogtsem_cmp Cgt v1 t1 v2 t2 m
  | Olesem_cmp Cle v1 t1 v2 t2 m
  | Ogesem_cmp Cge v1 t1 v2 t2 m
  end.

Definition sem_incrdecr (cenv: composite_env) (id: incr_or_decr) (v: val) (ty: type) (m: T) :=
  match id with
  | Incrsem_add cenv v ty (Vint Int.one) type_int32s m
  | Decrsem_sub cenv v ty (Vint Int.one) type_int32s m
  end.

Definition incrdecr_type (ty: type) :=
  match typeconv ty with
  | Tpointer ty aTpointer ty a
  | Tint sz sg aTint sz sg noattr
  | Tlong sg aTlong sg noattr
  | Tfloat sz aTfloat sz noattr
  | _Tvoid
  end.

End WITHSEMCAST.

Global Instance sem_cast_unit: SemCast (fun (_: unit) _ _false) := {}.

Global Instance sem_cast_mem `{memory_model_ops: Mem.MemoryModelOps}:
  SemCast Mem.valid_pointer
  := {}.

Section WITHMEMORYMODEL.
Context `{memory_model_prf: Mem.MemoryModel}.

Lemma weak_valid_pointer_eq:
  weak_valid_pointer = Mem.weak_valid_pointer.
Proof.
  reflexivity.
Qed.

Lemma sem_cast_unit_to_mem (m: mem) (u: unit) v t1 t2 v':
  sem_cast v t1 t2 u = Some v'
  sem_cast v t1 t2 m = Some v'.
Proof.
  unfold sem_cast.
  destruct (classify_cast t1 t2); auto.
  destruct v; auto.
  destruct Archi.ptr64; simpl;
  congruence.
Qed.

Lemma bool_val_unit_to_mem (m: mem) (u: unit) v t v':
  bool_val v t u = Some v'
  bool_val v t m = Some v'.
Proof.
  unfold bool_val.
  destruct (classify_bool t); auto.
  destruct v; auto.
  destruct Archi.ptr64; simpl; discriminate.
Qed.

Lemma sem_notbool_unit_to_mem (m: mem) (u: unit) v ty v':
  sem_notbool v ty u = Some v'
  sem_notbool v ty m = Some v'.
Proof.
  unfold sem_notbool.
  destruct (bool_val v ty u) eqn:?; simpl; try discriminate. intro A; inv A.
  erewrite bool_val_unit_to_mem; simpl; eauto.
Qed.

Lemma sem_binarith_unit_to_mem
      (m: mem) (u: unit)
      sem_int sem_long sem_float sem_single
      v1 t1 v2 t2
      v':
  sem_binarith sem_int sem_long sem_float sem_single v1 t1 v2 t2 u = Some v'
  sem_binarith sem_int sem_long sem_float sem_single v1 t1 v2 t2 m = Some v'.
Proof.
  unfold sem_binarith.
  destruct (sem_cast v1 t1 _ u) eqn:CAST1; [ | discriminate ].
  apply (sem_cast_unit_to_mem m) in CAST1.
  rewrite CAST1; clear CAST1.
  destruct (sem_cast v2 t2 _ u) eqn:CAST2; [ | discriminate ].
  apply (sem_cast_unit_to_mem m) in CAST2.
  rewrite CAST2; clear CAST2.
  auto.
Qed.

Lemma sem_add_unit_to_mem
      (m: mem) (u: unit)
      cenv v1 t1 v2 t2 v':
  sem_add cenv v1 t1 v2 t2 u = Some v'
  sem_add cenv v1 t1 v2 t2 m = Some v'.
Proof.
  unfold sem_add.
  destruct (classify_add t1 t2); auto.
  apply sem_binarith_unit_to_mem.
Qed.

Lemma sem_sub_unit_to_mem
      (m: mem) (u: unit)
      cenv v1 t1 v2 t2 v':
  sem_sub cenv v1 t1 v2 t2 u = Some v'
  sem_sub cenv v1 t1 v2 t2 m = Some v'.
Proof.
  unfold sem_sub.
  destruct (classify_sub t1 t2); auto.
  apply sem_binarith_unit_to_mem.
Qed.

Lemma sem_mul_unit_to_mem
      (m: mem) (u: unit)
      v1 t1 v2 t2 v':
  sem_mul v1 t1 v2 t2 u = Some v'
  sem_mul v1 t1 v2 t2 m = Some v'.
Proof.
  apply sem_binarith_unit_to_mem.
Qed.

Lemma sem_div_unit_to_mem
      (m: mem) (u: unit)
      v1 t1 v2 t2 v':
  sem_div v1 t1 v2 t2 u = Some v'
  sem_div v1 t1 v2 t2 m = Some v'.
Proof.
  apply sem_binarith_unit_to_mem.
Qed.

Lemma sem_mod_unit_to_mem
      (m: mem) (u: unit)
      v1 t1 v2 t2 v':
  sem_mod v1 t1 v2 t2 u = Some v'
  sem_mod v1 t1 v2 t2 m = Some v'.
Proof.
  apply sem_binarith_unit_to_mem.
Qed.

Lemma sem_and_unit_to_mem
      (m: mem) (u: unit)
      v1 t1 v2 t2 v':
  sem_and v1 t1 v2 t2 u = Some v'
  sem_and v1 t1 v2 t2 m = Some v'.
Proof.
  apply sem_binarith_unit_to_mem.
Qed.

Lemma sem_or_unit_to_mem
      (m: mem) (u: unit)
      v1 t1 v2 t2 v':
  sem_or v1 t1 v2 t2 u = Some v'
  sem_or v1 t1 v2 t2 m = Some v'.
Proof.
  apply sem_binarith_unit_to_mem.
Qed.

Lemma sem_xor_unit_to_mem
      (m: mem) (u: unit)
      v1 t1 v2 t2 v':
  sem_xor v1 t1 v2 t2 u = Some v'
  sem_xor v1 t1 v2 t2 m = Some v'.
Proof.
  apply sem_binarith_unit_to_mem.
Qed.

Lemma option_of_bool_cmpu_unit_to_mem
      (m: mem)
      c v1 v2 v':
  option_map Val.of_bool
             (Val.cmpu_bool (fun (_ : block) (_ : Z) ⇒ false) c v1 v2) =
  Some v'
  option_map Val.of_bool (Val.cmpu_bool (Mem.valid_pointer m) c v1 v2) =
  Some v'.
Proof.
  destruct c; destruct v1; destruct v2; simpl; try discriminate; auto;
  rewrite andb_comm; simpl; try discriminate;
  destruct (eq_block _ _); discriminate.
Qed.

Lemma option_of_bool_cmplu_unit_to_mem
      (m: mem)
      c v1 v2 v':
  option_map Val.of_bool
             (Val.cmplu_bool (fun (_ : block) (_ : Z) ⇒ false) c v1 v2) =
  Some v'
  option_map Val.of_bool (Val.cmplu_bool (Mem.valid_pointer m) c v1 v2) =
  Some v'.
Proof.
  destruct c; destruct v1; destruct v2; simpl; try discriminate; auto;
  rewrite andb_comm; simpl; try discriminate;
  destruct (eq_block _ _); discriminate.
Qed.

Lemma cmp_ptr_unit_to_mem:
   u m c v1 v2 v',
    cmp_ptr (T := unit) (valid_pointer := fun _ _ _false) u c v1 v2 = Some v'
    cmp_ptr (T := mem) (valid_pointer := Mem.valid_pointer) m c v1 v2 = Some v'.
Proof.
  unfold cmp_ptr. intros u m c v1 v2 v'.
  destruct Archi.ptr64.
  apply option_of_bool_cmplu_unit_to_mem.
  apply option_of_bool_cmpu_unit_to_mem.
Qed.

Lemma sem_cmp_unit_to_mem
      (m: mem) (u: unit)
      c
      v1 t1 v2 t2
      v':
  sem_cmp c v1 t1 v2 t2 u = Some v'
  sem_cmp c v1 t1 v2 t2 m = Some v'.
Proof.
  unfold sem_cmp.
  destruct (classify_cmp t1 t2).
  + apply cmp_ptr_unit_to_mem.
  + destruct v2; try congruence.
    apply cmp_ptr_unit_to_mem.
    destruct Archi.ptr64; auto; apply cmp_ptr_unit_to_mem.
  + destruct v1; auto; apply cmp_ptr_unit_to_mem.
  + destruct v2; auto; apply cmp_ptr_unit_to_mem.
  + destruct v1; auto; apply cmp_ptr_unit_to_mem.
  + apply sem_binarith_unit_to_mem.
Qed.

Lemma sem_unary_operation_unit_to_mem
      (m: mem) (u: unit)
      op v ty v':
  sem_unary_operation op v ty u = Some v'
  sem_unary_operation op v ty m = Some v'.
Proof.
  unfold sem_unary_operation.
  destruct op; auto.
  apply sem_notbool_unit_to_mem.
Qed.

Lemma sem_binary_operation_unit_to_mem
      (m: mem) (u: unit)
      cenv op v1 t1 v2 t2 v':
  sem_binary_operation cenv op v1 t1 v2 t2 u = Some v'
  sem_binary_operation cenv op v1 t1 v2 t2 m = Some v'.
Proof.
  unfold sem_binary_operation.
  destruct op; eauto using sem_cmp_unit_to_mem.
  + apply sem_add_unit_to_mem.
  + apply sem_sub_unit_to_mem.
  + apply sem_mul_unit_to_mem.
  + apply sem_div_unit_to_mem.
  + apply sem_mod_unit_to_mem.
  + apply sem_and_unit_to_mem.
  + apply sem_or_unit_to_mem.
  + apply sem_xor_unit_to_mem.
Qed.

Lemma sem_incrdecr_unit_to_mem
      (m: mem) (u: unit)
      cenv id v ty v':
  sem_incrdecr cenv id v ty u = Some v'
  sem_incrdecr cenv id v ty m = Some v'.
Proof.
  destruct id.
  + apply sem_add_unit_to_mem.
  + apply sem_sub_unit_to_mem.
Qed.

Compatibility with extensions and injections


Section GENERIC_INJECTION.

Variable f: meminj.
Variables m m': mem.

Hypothesis valid_pointer_inj:
   b1 ofs b2 delta,
  f b1 = Some(b2, delta)
  Mem.valid_pointer m b1 (Ptrofs.unsigned ofs) = true
  Mem.valid_pointer m' b2 (Ptrofs.unsigned (Ptrofs.add ofs (Ptrofs.repr delta))) = true.

Hypothesis weak_valid_pointer_inj:
   b1 ofs b2 delta,
  f b1 = Some(b2, delta)
  Mem.weak_valid_pointer m b1 (Ptrofs.unsigned ofs) = true
  Mem.weak_valid_pointer m' b2 (Ptrofs.unsigned (Ptrofs.add ofs (Ptrofs.repr delta))) = true.

Hypothesis weak_valid_pointer_no_overflow:
   b1 ofs b2 delta,
  f b1 = Some(b2, delta)
  Mem.weak_valid_pointer m b1 (Ptrofs.unsigned ofs) = true
  0 Ptrofs.unsigned ofs + Ptrofs.unsigned (Ptrofs.repr delta) Ptrofs.max_unsigned.

Hypothesis valid_different_pointers_inj:
   b1 ofs1 b2 ofs2 b1' delta1 b2' delta2,
  b1 b2
  Mem.valid_pointer m b1 (Ptrofs.unsigned ofs1) = true
  Mem.valid_pointer m b2 (Ptrofs.unsigned ofs2) = true
  f b1 = Some (b1', delta1)
  f b2 = Some (b2', delta2)
  b1' b2'
  Ptrofs.unsigned (Ptrofs.add ofs1 (Ptrofs.repr delta1)) Ptrofs.unsigned (Ptrofs.add ofs2 (Ptrofs.repr delta2)).

Remark val_inject_vtrue: f, Val.inject f Vtrue Vtrue.
Proof. unfold Vtrue; auto. Qed.

Remark val_inject_vfalse: f, Val.inject f Vfalse Vfalse.
Proof. unfold Vfalse; auto. Qed.

Remark val_inject_of_bool: f b, Val.inject f (Val.of_bool b) (Val.of_bool b).
Proof. intros. unfold Val.of_bool. destruct b; [apply val_inject_vtrue|apply val_inject_vfalse].
Qed.

Remark val_inject_vptrofs: n, Val.inject f (Vptrofs n) (Vptrofs n).
Proof. intros. unfold Vptrofs. destruct Archi.ptr64; auto. Qed.

Hint Resolve val_inject_vtrue val_inject_vfalse val_inject_of_bool val_inject_vptrofs.

Ltac TrivialInject :=
  match goal with
  | [ H: None = Some _ |- _ ] ⇒ discriminate
  | [ H: Some _ = Some _ |- _ ] ⇒ inv H; TrivialInject
  | [ H: match ?x with Some __ | None_ end = Some _ |- _ ] ⇒ destruct x; TrivialInject
  | [ H: match ?x with true_ | false_ end = Some _ |- _ ] ⇒ destruct x eqn:?; TrivialInject
  | [ |- v', Some ?v = Some v' _ ] ⇒ v; split; auto
  | _idtac
  end.

Lemma sem_cast_inj:
   v1 ty1 ty v tv1,
  sem_cast v1 ty1 ty m = Some v
  Val.inject f v1 tv1
   tv, sem_cast tv1 ty1 ty m'= Some tv Val.inject f v tv.
Proof.
  unfold sem_cast; intros v1 ty1 ty v tv1 SC VINJ;
    destruct (classify_cast ty1 ty); inv VINJ; TrivialInject.
  - econstructor; eauto.
  - rewrite weak_valid_pointer_eq in × |- ×.
    erewrite weak_valid_pointer_inj by eauto. TrivialInject.
  - rewrite weak_valid_pointer_eq in × |- ×.
    erewrite weak_valid_pointer_inj by eauto. TrivialInject.
  - destruct (ident_eq id1 id2); TrivialInject. econstructor; eauto.
  - destruct (ident_eq id1 id2); TrivialInject. econstructor; eauto.
  - econstructor; eauto.
Qed.

Lemma bool_val_inj:
   v ty b tv,
  bool_val v ty m = Some b
  Val.inject f v tv
  bool_val tv ty m' = Some b.
Proof.
  unfold bool_val; intros.
  rewrite weak_valid_pointer_eq in × |- ×.
  destruct (classify_bool ty); inv H0; try congruence.
  destruct Archi.ptr64; try discriminate.
  destruct (Mem.weak_valid_pointer m b1 (Ptrofs.unsigned ofs1)) eqn:VP; inv H.
  erewrite weak_valid_pointer_inj by eauto. auto.
  destruct Archi.ptr64; try discriminate.
  destruct (Mem.weak_valid_pointer m b1 (Ptrofs.unsigned ofs1)) eqn:VP; inv H.
  erewrite weak_valid_pointer_inj by eauto. auto.
Qed.

Lemma sem_unary_operation_inj:
   op v1 ty v tv1,
  sem_unary_operation op v1 ty m = Some v
  Val.inject f v1 tv1
   tv, sem_unary_operation op tv1 ty m' = Some tv Val.inject f v tv.
Proof.
  unfold sem_unary_operation; intros.
  destruct op; try rewrite weak_valid_pointer_eq in × |- ×.
-
  unfold sem_notbool in ×. destruct (bool_val v1 ty m) as [b|] eqn:BV; simpl in H; inv H.
  erewrite bool_val_inj by eauto. simpl. TrivialInject.
-
  unfold sem_notint in *; destruct (classify_notint ty); inv H0; inv H; TrivialInject.
-
  unfold sem_neg in *; destruct (classify_neg ty); inv H0; inv H; TrivialInject.
-
  unfold sem_absfloat in *; destruct (classify_neg ty); inv H0; inv H; TrivialInject.
Qed.

Definition optval_self_injects (ov: option val) : Prop :=
  match ov with
  | Some (Vptr b ofs) ⇒ False
  | _True
  end.

Remark sem_binarith_inject:
   sem_int sem_long sem_float sem_single v1 t1 v2 t2 v v1' v2',
  sem_binarith sem_int sem_long sem_float sem_single v1 t1 v2 t2 m = Some v
  Val.inject f v1 v1' Val.inject f v2 v2'
  ( sg n1 n2, optval_self_injects (sem_int sg n1 n2))
  ( sg n1 n2, optval_self_injects (sem_long sg n1 n2))
  ( n1 n2, optval_self_injects (sem_float n1 n2))
  ( n1 n2, optval_self_injects (sem_single n1 n2))
   v', sem_binarith sem_int sem_long sem_float sem_single v1' t1 v2' t2 m' = Some v' Val.inject f v v'.
Proof.
  intros.
  assert (SELF: ov v, ov = Some v optval_self_injects ov Val.inject f v v).
  {
    intros. subst ov; simpl in H7. destruct v0; contradiction || constructor.
  }
  unfold sem_binarith in ×.
  set (c := classify_binarith t1 t2) in ×.
  set (t := binarith_type c) in ×.
  destruct (sem_cast v1 t1 t m) as [w1|] eqn:C1; try discriminate.
  destruct (sem_cast v2 t2 t m) as [w2|] eqn:C2; try discriminate.
  exploit (sem_cast_inj v1); eauto. intros (w1' & C1' & INJ1).
  exploit (sem_cast_inj v2); eauto. intros (w2' & C2' & INJ2).
  rewrite C1'; rewrite C2'.
  destruct c; inv INJ1; inv INJ2; discriminate || eauto.
Qed.

Remark sem_shift_inject:
   sem_int sem_long v1 t1 v2 t2 v v1' v2',
  sem_shift sem_int sem_long v1 t1 v2 t2 = Some v
  Val.inject f v1 v1' Val.inject f v2 v2'
   v', sem_shift sem_int sem_long v1' t1 v2' t2 = Some v' Val.inject f v v'.
Proof.
  intros. v.
  unfold sem_shift in *; destruct (classify_shift t1 t2); inv H0; inv H1; try discriminate.
  destruct (Int.ltu i0 Int.iwordsize); inv H; auto.
  destruct (Int64.ltu i0 Int64.iwordsize); inv H; auto.
  destruct (Int64.ltu i0 (Int64.repr 32)); inv H; auto.
  destruct (Int.ltu i0 Int64.iwordsize'); inv H; auto.
Qed.

Remark sem_cmp_ptr_inj:
   c v1 v2 v tv1 tv2,
  cmp_ptr (valid_pointer := Mem.valid_pointer) m c v1 v2 = Some v
  Val.inject f v1 tv1
  Val.inject f v2 tv2
   tv, cmp_ptr (valid_pointer := Mem.valid_pointer) m' c tv1 tv2 = Some tv Val.inject f v tv.
Proof.
  unfold cmp_ptr; intros.
  remember (if Archi.ptr64
       then Val.cmplu_bool (Mem.valid_pointer m) c v1 v2
       else Val.cmpu_bool (Mem.valid_pointer m) c v1 v2) as ob.
  destruct ob as [b|]; simpl in H; inv H.
   (Val.of_bool b); split; auto.
  destruct Archi.ptr64.
  erewrite Val.cmplu_bool_inject by eauto. auto.
  erewrite Val.cmpu_bool_inject by eauto. auto.
Qed.

Remark sem_cmp_inj:
   cmp v1 tv1 ty1 v2 tv2 ty2 v,
  sem_cmp cmp v1 ty1 v2 ty2 m = Some v
  Val.inject f v1 tv1
  Val.inject f v2 tv2
   tv, sem_cmp cmp tv1 ty1 tv2 ty2 m' = Some tv Val.inject f v tv.
Proof.
  intros.
  unfold sem_cmp in *; destruct (classify_cmp ty1 ty2).
-
  eapply sem_cmp_ptr_inj; eauto.
-
  inversion H1; subst; TrivialInject; eapply sem_cmp_ptr_inj; eauto.
-
  inversion H0; subst; TrivialInject; eapply sem_cmp_ptr_inj; eauto.
-
  inversion H1; subst; TrivialInject; eapply sem_cmp_ptr_inj; eauto.
-
  inversion H0; subst; TrivialInject; eapply sem_cmp_ptr_inj; eauto.
-
  assert (SELF: b, optval_self_injects (Some (Val.of_bool b))).
  {
    destruct b; exact I.
  }
  eapply sem_binarith_inject; eauto.
Qed.

Lemma sem_binary_operation_inj:
   cenv op v1 ty1 v2 ty2 v tv1 tv2,
  sem_binary_operation cenv op v1 ty1 v2 ty2 m = Some v
  Val.inject f v1 tv1 Val.inject f v2 tv2
   tv, sem_binary_operation cenv op tv1 ty1 tv2 ty2 m' = Some tv Val.inject f v tv.
Proof.
  unfold sem_binary_operation; intros; destruct op.
-
  assert (A: cenv ty si v1' v2' tv1' tv2',
             Val.inject f v1' tv1' Val.inject f v2' tv2'
             sem_add_ptr_int cenv ty si v1' v2' = Some v
              tv, sem_add_ptr_int cenv ty si tv1' tv2' = Some tv Val.inject f v tv).
  { intros. unfold sem_add_ptr_int in *; inv H2; inv H3; TrivialInject.
    econstructor. eauto. repeat rewrite Ptrofs.add_assoc. decEq. apply Ptrofs.add_commut. }
  assert (B: cenv ty v1' v2' tv1' tv2',
             Val.inject f v1' tv1' Val.inject f v2' tv2'
             sem_add_ptr_long cenv ty v1' v2' = Some v
              tv, sem_add_ptr_long cenv ty tv1' tv2' = Some tv Val.inject f v tv).
  { intros. unfold sem_add_ptr_long in *; inv H2; inv H3; TrivialInject.
    econstructor. eauto. repeat rewrite Ptrofs.add_assoc. decEq. apply Ptrofs.add_commut. }
  unfold sem_add in *; destruct (classify_add ty1 ty2); eauto.
  + eapply sem_binarith_inject; eauto; intros; exact I.
-
  unfold sem_sub in *; destruct (classify_sub ty1 ty2).
  + inv H0; inv H1; TrivialInject.
    econstructor. eauto. rewrite Ptrofs.sub_add_l. auto.
  + inv H0; inv H1; TrivialInject.
    destruct (eq_block b1 b0); try discriminate. subst b1.
    rewrite H0 in H2; inv H2. rewrite dec_eq_true.
    destruct (zlt 0 (sizeof cenv ty) && zle (sizeof cenv ty) Ptrofs.max_signed); inv H.
    rewrite Ptrofs.sub_shifted. TrivialInject.
  + inv H0; inv H1; TrivialInject.
    econstructor. eauto. rewrite Ptrofs.sub_add_l. auto.
  + eapply sem_binarith_inject; eauto; intros; exact I.
-
  eapply sem_binarith_inject; eauto; intros; exact I.
-
  eapply sem_binarith_inject; eauto; intros.
  destruct sg.
  destruct (Int.eq n2 Int.zero
            || Int.eq n1 (Int.repr Int.min_signed) && Int.eq n2 Int.mone); exact I.
  destruct (Int.eq n2 Int.zero); exact I.
  destruct sg.
  destruct (Int64.eq n2 Int64.zero
            || Int64.eq n1 (Int64.repr Int64.min_signed) && Int64.eq n2 Int64.mone); exact I.
  destruct (Int64.eq n2 Int64.zero); exact I.
  exact I.
  exact I.
-
  eapply sem_binarith_inject; eauto; intros.
  destruct sg.
  destruct (Int.eq n2 Int.zero
            || Int.eq n1 (Int.repr Int.min_signed) && Int.eq n2 Int.mone); exact I.
  destruct (Int.eq n2 Int.zero); exact I.
  destruct sg.
  destruct (Int64.eq n2 Int64.zero
            || Int64.eq n1 (Int64.repr Int64.min_signed) && Int64.eq n2 Int64.mone); exact I.
  destruct (Int64.eq n2 Int64.zero); exact I.
  exact I.
  exact I.
-
  eapply sem_binarith_inject; eauto; intros; exact I.
-
  eapply sem_binarith_inject; eauto; intros; exact I.
-
  eapply sem_binarith_inject; eauto; intros; exact I.
-
  eapply sem_shift_inject; eauto.
-
  eapply sem_shift_inject; eauto.
- eapply sem_cmp_inj; eauto.
- eapply sem_cmp_inj; eauto.
- eapply sem_cmp_inj; eauto.
- eapply sem_cmp_inj; eauto.
- eapply sem_cmp_inj; eauto.
- eapply sem_cmp_inj; eauto.
Qed.

End GENERIC_INJECTION.

Lemma sem_cast_inject:
   f v1 ty1 ty m v tv1 tm,
  sem_cast v1 ty1 ty m = Some v
  Val.inject f v1 tv1
  Mem.inject f m tm
   tv, sem_cast tv1 ty1 ty tm = Some tv Val.inject f v tv.
Proof.
  intros. eapply sem_cast_inj; eauto.
  intros; eapply Mem.weak_valid_pointer_inject_val; eauto.
Qed.

Lemma sem_unary_operation_inject:
   f m m' op v1 ty1 v tv1,
  sem_unary_operation op v1 ty1 m = Some v
  Val.inject f v1 tv1
  Mem.inject f m m'
   tv, sem_unary_operation op tv1 ty1 m' = Some tv Val.inject f v tv.
Proof.
  intros. eapply sem_unary_operation_inj; eauto.
  intros; eapply Mem.weak_valid_pointer_inject_val; eauto.
Qed.

Lemma sem_binary_operation_inject:
   f m m' cenv op v1 ty1 v2 ty2 v tv1 tv2,
  sem_binary_operation cenv op v1 ty1 v2 ty2 m = Some v
  Val.inject f v1 tv1 Val.inject f v2 tv2
  Mem.inject f m m'
   tv, sem_binary_operation cenv op tv1 ty1 tv2 ty2 m' = Some tv Val.inject f v tv.
Proof.
  intros. eapply sem_binary_operation_inj; eauto.
  intros; eapply Mem.valid_pointer_inject_val; eauto.
  intros; eapply Mem.weak_valid_pointer_inject_val; eauto.
  intros; eapply Mem.weak_valid_pointer_inject_no_overflow; eauto.
  intros; eapply Mem.different_pointers_inject; eauto.
Qed.

Lemma bool_val_inject:
   f m m' v ty b tv,
  bool_val v ty m = Some b
  Val.inject f v tv
  Mem.inject f m m'
  bool_val tv ty m' = Some b.
Proof.
  intros. eapply bool_val_inj; eauto.
  intros; eapply Mem.weak_valid_pointer_inject_val; eauto.
Qed.

Some properties of operator semantics

This section collects some common-sense properties about the type classification and semantic functions above. Some properties are used in the CompCert semantics preservation proofs. Others are not, but increase confidence in the specification and its relation with the ISO C99 standard.
Relation between Boolean value and casting to _Bool type.

Lemma cast_bool_bool_val:
   v t m,
  sem_cast v t (Tint IBool Signed noattr) m =
  match bool_val v t m with NoneNone | Some bSome(Val.of_bool b) end.
  intros.
  assert (A: classify_bool t =
    match t with
    | Tint _ _ _bool_case_i
    | Tpointer _ _ | Tarray _ _ _ | Tfunction _ _ _if Archi.ptr64 then bool_case_l else bool_case_i
    | Tfloat F64 _bool_case_f
    | Tfloat F32 _bool_case_s
    | Tlong _ _bool_case_l
    | _bool_default
    end).
  {
    unfold classify_bool; destruct t; simpl; auto. destruct i; auto.
  }
  unfold bool_val. rewrite A.
  unfold sem_cast, classify_cast; remember Archi.ptr64 as ptr64; destruct t; simpl; auto; destruct v; auto;
    try rewrite weak_valid_pointer_eq in × |- ×.
  destruct (Int.eq i0 Int.zero); auto.
  destruct ptr64; auto. destruct (Mem.weak_valid_pointer m b (Ptrofs.unsigned i0)); auto.
  destruct (Int64.eq i Int64.zero); auto.
  destruct (negb ptr64); auto. destruct (Mem.weak_valid_pointer m b (Ptrofs.unsigned i)); auto.
  destruct f; auto.
  destruct f; auto.
  destruct f; auto.
  destruct f; auto.
  destruct (Float.cmp Ceq f0 Float.zero); auto.
  destruct f; auto.
  destruct (Float32.cmp Ceq f0 Float32.zero); auto.
  destruct f; auto.
  destruct ptr64; auto.
  destruct (Int.eq i Int.zero); auto.
  destruct ptr64; auto.
  destruct ptr64; auto.
  destruct ptr64; auto. destruct (Int64.eq i Int64.zero); auto.
  destruct ptr64; auto.
  destruct ptr64; auto.
  destruct ptr64; auto. destruct (Mem.weak_valid_pointer m b (Ptrofs.unsigned i)); auto.
  destruct (Mem.weak_valid_pointer m b (Ptrofs.unsigned i)); auto.
  destruct ptr64; auto.
  destruct ptr64; auto. destruct (Int.eq i Int.zero); auto.
  destruct ptr64; auto. destruct (Int64.eq i Int64.zero); auto.
  destruct ptr64; auto.
  destruct ptr64; auto.
  destruct ptr64; auto. destruct (Mem.weak_valid_pointer m b (Ptrofs.unsigned i)); auto.
  destruct (Mem.weak_valid_pointer m b (Ptrofs.unsigned i)); auto.
  destruct ptr64; auto.
  destruct ptr64; auto. destruct (Int.eq i Int.zero); auto.
  destruct ptr64; auto. destruct (Int64.eq i Int64.zero); auto.
  destruct ptr64; auto.
  destruct ptr64; auto.
  destruct ptr64; auto. destruct (Mem.weak_valid_pointer m b (Ptrofs.unsigned i)); auto.
  destruct (Mem.weak_valid_pointer m b (Ptrofs.unsigned i)); auto.
Qed.

Relation between Boolean value and Boolean negation.

Lemma notbool_bool_val:
   v t m,
  sem_notbool v t m =
  match bool_val v t m with NoneNone | Some bSome(Val.of_bool (negb b)) end.
Proof.
  intros. unfold sem_notbool. destruct (bool_val v t m) as [[] | ]; reflexivity.
Qed.

Properties of values obtained by casting to a given type.

Section VAL_CASTED.

Inductive val_casted: val type Prop :=
  | val_casted_int: sz si attr n,
      cast_int_int sz si n = n
      val_casted (Vint n) (Tint sz si attr)
  | val_casted_float: attr n,
       val_casted (Vfloat n) (Tfloat F64 attr)
  | val_casted_single: attr n,
       val_casted (Vsingle n) (Tfloat F32 attr)
  | val_casted_long: si attr n,
      val_casted (Vlong n) (Tlong si attr)
  | val_casted_ptr_ptr: b ofs ty attr,
      val_casted (Vptr b ofs) (Tpointer ty attr)
  | val_casted_int_ptr: n ty attr,
      Archi.ptr64 = false val_casted (Vint n) (Tpointer ty attr)
  | val_casted_ptr_int: b ofs si attr,
      Archi.ptr64 = false val_casted (Vptr b ofs) (Tint I32 si attr)
  | val_casted_long_ptr: n ty attr,
      Archi.ptr64 = true val_casted (Vlong n) (Tpointer ty attr)
  | val_casted_ptr_long: b ofs si attr,
      Archi.ptr64 = true val_casted (Vptr b ofs) (Tlong si attr)
  | val_casted_struct: id attr b ofs,
      val_casted (Vptr b ofs) (Tstruct id attr)
  | val_casted_union: id attr b ofs,
      val_casted (Vptr b ofs) (Tunion id attr)
  | val_casted_void: v,
      val_casted v Tvoid.

Hint Constructors val_casted.

Remark cast_int_int_idem:
   sz sg i, cast_int_int sz sg (cast_int_int sz sg i) = cast_int_int sz sg i.
Proof.
  intros. destruct sz; simpl; auto.
  destruct sg; [apply Int.sign_ext_idem|apply Int.zero_ext_idem]; compute; intuition congruence.
  destruct sg; [apply Int.sign_ext_idem|apply Int.zero_ext_idem]; compute; intuition congruence.
  destruct (Int.eq i Int.zero); auto.
Qed.

Ltac DestructCases :=
  match goal with
  | [H: match match ?x with __ end with __ end = Some _ |- _ ] ⇒ destruct x eqn:?; DestructCases
  | [H: match ?x with __ end = Some _ |- _ ] ⇒ destruct x eqn:?; DestructCases
  | [H: Some _ = Some _ |- _ ] ⇒ inv H; DestructCases
  | [H: None = Some _ |- _ ] ⇒ discriminate H
  | [H: @eq intsize _ _ |- _ ] ⇒ discriminate H || (clear H; DestructCases)
  | [ |- val_casted (Vint (if ?x then Int.zero else Int.one)) _ ] ⇒
       try (constructor; destruct x; reflexivity)
  | [ |- val_casted (Vint _) (Tint ?sz ?sg _) ] ⇒
       try (constructor; apply (cast_int_int_idem sz sg))
  | _idtac
  end.

Lemma cast_val_is_casted:
   v ty ty' v' m, sem_cast v ty ty' m = Some v' val_casted v' ty'.
Proof.
  unfold sem_cast; intros.
  destruct ty, ty'; simpl in H; DestructCases; constructor; auto.
Qed.

End VAL_CASTED.

As a consequence, casting twice is equivalent to casting once.

Lemma cast_val_casted:
   v ty m, val_casted v ty sem_cast v ty ty m = Some v.
Proof.
  intros. unfold sem_cast; inversion H; clear H; subst v ty; simpl.
- destruct Archi.ptr64; [ | destruct (intsize_eq sz I32)].
+ destruct sz; f_equal; f_equal; assumption.
+ subst sz; auto.
+ destruct sz; f_equal; f_equal; assumption.
- auto.
- auto.
- destruct Archi.ptr64; auto.
- auto.
- rewrite H0; auto.
- rewrite H0; auto.
- rewrite H0; auto.
- rewrite H0; auto.
- rewrite dec_eq_true; auto.
- rewrite dec_eq_true; auto.
- auto.
Qed.

Lemma cast_idempotent:
   v ty ty' v' m, sem_cast v ty ty' m = Some v' sem_cast v' ty' ty' m = Some v'.
Proof.
  intros. apply cast_val_casted. eapply cast_val_is_casted; eauto.
Qed.

End WITHMEMORYMODEL.

Relation with the arithmetic conversions of ISO C99, section 6.3.1

Module ArithConv.

This is the ISO C algebra of arithmetic types, without qualifiers. S stands for "signed" and U for "unsigned".

Inductive int_type : Type :=
  | _Bool
  | Char | SChar | UChar
  | Short | UShort
  | Int | UInt
  | Long | ULong
  | Longlong | ULonglong.

Inductive arith_type : Type :=
  | I (it: int_type)
  | Float
  | Double
  | Longdouble.

Definition eq_int_type: (x y: int_type), {x=y} + {xy}.
Proof. decide equality. Defined.

Definition is_unsigned (t: int_type) : bool :=
  match t with
  | _Booltrue
  | Charfalse
  | SCharfalse
  | UChartrue
  | Shortfalse
  | UShorttrue
  | Intfalse
  | UInttrue
  | Longfalse
  | ULongtrue
  | Longlongfalse
  | ULonglongtrue
  end.

Definition unsigned_type (t: int_type) : int_type :=
  match t with
  | CharUChar
  | SCharUChar
  | ShortUShort
  | IntUInt
  | LongULong
  | LonglongULonglong
  | _t
  end.

Definition int_sizeof (t: int_type) : Z :=
  match t with
  | _Bool | Char | SChar | UChar ⇒ 1
  | Short | UShort ⇒ 2
  | Int | UInt | Long | ULong ⇒ 4
  | Longlong | ULonglong ⇒ 8
  end.

6.3.1.1 para 1: integer conversion rank

Definition rank (t: int_type) : Z :=
  match t with
  | _Bool ⇒ 1
  | Char | SChar | UChar ⇒ 2
  | Short | UShort ⇒ 3
  | Int | UInt ⇒ 4
  | Long | ULong ⇒ 5
  | Longlong | ULonglong ⇒ 6
  end.

6.3.1.1 para 2: integer promotions, a.k.a. usual unary conversions

Definition integer_promotion (t: int_type) : int_type :=
  if zlt (rank t) (rank Int) then Int else t.

6.3.1.8: Usual arithmetic conversions, a.k.a. binary conversions. This function returns the type to which the two operands must be converted.

Definition usual_arithmetic_conversion (t1 t2: arith_type) : arith_type :=
  match t1, t2 with
  
  | Longdouble, _ | _, LongdoubleLongdouble
  
  | Double, _ | _, DoubleDouble
  
  | Float, _ | _, FloatFloat
  
  | I i1, I i2
    let j1 := integer_promotion i1 in
    let j2 := integer_promotion i2 in
    
    if eq_int_type j1 j2 then I j1 else
    match is_unsigned j1, is_unsigned j2 with
    
    | true, true | false, false
        if zlt (rank j1) (rank j2) then I j2 else I j1
    | true, false
    
        if zle (rank j2) (rank j1) then I j1 else
    
        if zlt (int_sizeof j1) (int_sizeof j2) then I j2 else
    
        I (unsigned_type j2)
    | false, true
    
        if zle (rank j1) (rank j2) then I j2 else
        if zlt (int_sizeof j2) (int_sizeof j1) then I j1 else
        I (unsigned_type j1)
    end
  end.

Mapping ISO arithmetic types to CompCert types

Definition proj_type (t: arith_type) : type :=
  match t with
  | I _BoolTint IBool Unsigned noattr
  | I CharTint I8 Unsigned noattr
  | I SCharTint I8 Signed noattr
  | I UCharTint I8 Unsigned noattr
  | I ShortTint I16 Signed noattr
  | I UShortTint I16 Unsigned noattr
  | I IntTint I32 Signed noattr
  | I UIntTint I32 Unsigned noattr
  | I LongTint I32 Signed noattr
  | I ULongTint I32 Unsigned noattr
  | I LonglongTlong Signed noattr
  | I ULonglongTlong Unsigned noattr
  | FloatTfloat F32 noattr
  | DoubleTfloat F64 noattr
  | LongdoubleTfloat F64 noattr
  end.

Relation between typeconv and integer promotion.

Lemma typeconv_integer_promotion:
   i, typeconv (proj_type (I i)) = proj_type (I (integer_promotion i)).
Proof.
  destruct i; reflexivity.
Qed.

Relation between classify_binarith and arithmetic conversion.

Lemma classify_binarith_arithmetic_conversion:
   t1 t2,
  binarith_type (classify_binarith (proj_type t1) (proj_type t2)) =
  proj_type (usual_arithmetic_conversion t1 t2).
Proof.
  destruct t1; destruct t2; try reflexivity.
- destruct it; destruct it0; reflexivity.
- destruct it; reflexivity.
- destruct it; reflexivity.
- destruct it; reflexivity.
- destruct it; reflexivity.
- destruct it; reflexivity.
- destruct it; reflexivity.
Qed.

End ArithConv.