let rec lval_to_loc ~with_alarms state lv =
  let _,_,r =
    lval_to_loc_deps_option
      ~with_alarms
      ~deps:None
      ~reduce_valid_index:(Parameters.SafeArrays.get ())
      state
      lv
  in
  r

and lval_to_loc_deps_option
    ~with_alarms ~deps (state:Relations_type.Model.t) ~reduce_valid_index
    (base,offset as lv)  =
  if not (Relations_type.Model.is_reachable state) then
    state, deps, loc_bottom
  else
    let typ = match base with
    | Var host -> host.vtype
    | Mem x -> typeOf x
    in
    try
      let state, deps, offs =
        eval_offset
          ~reduce_valid_index
          ~with_alarms deps typ state offset
      in
      base_to_loc ~with_alarms ?deps state lv base offs
    with Offset_not_based_on_Null(deps,offset) ->
      let state, deps, loc_if_there_wasnt_offset =
        base_to_loc ~with_alarms ?deps state lv base Ival.zero
      in
      state, deps,
      loc_bits_to_loc lv
        (Location_Bits.join
            (loc_bytes_to_loc_bits offset)
            loc_if_there_wasnt_offset.loc)

(* pc says: only called in addrOf *)
and lval_to_loc_with_offset_deps_only
    ~deps (state:Relations_type.Model.t) v
    =
  lval_to_loc_with_offset_deps_only_option ~deps:(Some deps) state v

and lval_to_loc_with_deps ~deps state lv ~with_alarms =
  lval_to_loc_deps_option ~with_alarms ~deps:(Some deps) state lv

(* pc says: only called in addrOf *)
and lval_to_loc_with_offset_deps_only_option
    ~with_alarms ~deps (state:Relations_type.Model.t) (_base, _offset as v)
    =
  lval_to_loc_deps_option
    ~with_alarms ~deps (state:Relations_type.Model.t) (v)
    ~reduce_valid_index:false


(** Detects if an expression can be considered as a lvalue even though it is hidden by a cast that does not change the lvalue. Raises exn if it is not an lvalue.

TODO: When the goal is to recognize the form (cast)l-value == expr, it would be better and more powerful to have chains of inverse functions *)



and pass_cast ~with_alarms state exn typ e =
  (* type might be the same but their attributes.
     But volatile attribute cannot be skipped *)

  if not (Cilutil.equals
             (typeSigWithAttrs (filterAttributes "volatile") typ)
             (typeSigWithAttrs (filterAttributes "volatile") (typeOf e)))
  then
    (try
        let typeofe = typeOf e in
        (* Any volatile attribute may have an effect on the expression value *)
        if hasAttribute "volatile" (typeAttrs typeofe)
          || hasAttribute  "volatile" (typeAttrs typ)
        then raise exn;
        let sztyp = sizeof typ in
        let szexpr = sizeof typeofe in
        let typ_ge_typeofe =
          match sztyp,szexpr with
            Int_Base.Value styp, Int_Base.Value sexpr -> Int.ge styp sexpr
          | _ -> false
        in
        if typ_ge_typeofe then
          let sityp = is_signed_int_enum_pointer typ in
          let sisexpr = is_signed_int_enum_pointer (typeOf e) in
          if sityp = sisexpr then ()
            (* destination type is larger and has the same sign as
               the original type *)

          else begin (* try to ignore the cast if it acts as identity
                        on the value [e] even if signed/unsigned
                        conflict. *)

              match unrollType typ with
              | TInt _ | TEnum _ ->
                  let size = Int.of_int (bitsSizeOf typ) in
                  let signed = sityp in
                  (try
                      let old_ival = V.find_ival
                        (eval_expr ~with_alarms state e)
                      in
                      if (Ival.equal
                             old_ival
                             (Ival.cast ~size ~signed ~value:old_ival))
                      then () (* [e] is not sensitive to cast *)
                      else raise exn
                    with
                    | Not_found
                    | V.Not_based_on_null ->
                        raise exn)
                    (* this is not always injective, thus cannot be
                       easily reverted. *)

              | _ -> raise exn
            end
        else raise exn
      with Neither_Int_Nor_Enum_Nor_Pointer
        -> raise exn)

and find_lv ~with_alarms (state:Relations_type.Model.t) ee =
  (* [BM] Do not recognize an lval whenever a volatile is involved to
     prevent copy/paste optimization. IS THIS THE RIGHTPLACE PC ?*)

  if hasAttribute "volatile" (typeAttrs (typeOf ee)) then
    raise Cannot_find_lv;
  match ee.enode with
  | Lval lv -> lv
  | CastE (typ,e) ->
      ( match unrollType typ, unrollType (typeOf e) with
        TFloat _, TFloat _ -> find_lv ~with_alarms state e
          (* see remark at pass_cast about inverse functions *)
      | _ ->
          pass_cast ~with_alarms state Cannot_find_lv typ e;
          find_lv ~with_alarms state e)
  | _ -> raise Cannot_find_lv

and find_lv_plus ~with_alarms state e =
  let acc = ref [] in
  let rec find_lv_plus_rec e current_offs =
    try
      let lv = find_lv ~with_alarms state e in
      if not (hasAttribute "volatile" (typeAttrs (Cil.typeOfLval lv)))
      then acc := (lv,current_offs) :: !acc
    with Cannot_find_lv ->
      match e.enode with
        BinOp(op, e1, e2, typ) ->
          begin
            match unrollType typ with
              TFloat _ -> ()
            | _ -> begin
                  match op with
                    PlusA ->
                      let ev1 = eval_expr ~with_alarms state e1 in
                      let ev2 = eval_expr ~with_alarms state e2 in
                      ( try
                          let ival1 = V.find_ival ev1 in
                          find_lv_plus_rec e2 (Ival.add current_offs ival1)
                        with V.Not_based_on_null -> ());
                      ( try
                          let ival2 = V.find_ival ev2 in
                          find_lv_plus_rec e1 (Ival.add current_offs ival2)
                        with V.Not_based_on_null -> ());
                  | (MinusA|MinusPI|PlusPI|IndexPI as b) ->
                      let ev2 = eval_expr ~with_alarms state e2 in
                      ( try
                          let ival2 = V.find_ival ev2 in
                          let ival2 =
                            if b = MinusA
                            then ival2
                            else
                              let ival2 =
                                Ival.scale
                                  (Int_Base.project (osizeof_pointed typ))
                                  ival2
                              in
                              if b = MinusPI
                              then ival2
                              else Ival.neg ival2
                          in
                          find_lv_plus_rec e1 (Ival.sub current_offs ival2)
                        with V.Not_based_on_null | Int_Base.Error_Top-> ());
                  | _ -> ()
              end
          end
      | CastE(typ,e) ->
          ( try
              pass_cast ~with_alarms  state Cannot_find_lv typ e;
            find_lv_plus_rec e current_offs
            with Cannot_find_lv -> ())
      | _ -> ()
  in
  find_lv_plus_rec e Ival.singleton_zero;
  (*List.iter
    (fun (lv,ival) ->
    ignore (Pretty.printf "find_lv_plus %a : %s\n"
    d_lval lv
    (pretty_to_string Ival.pretty ival)))
    !acc;*)

  !acc

and base_to_loc ~with_alarms ?deps state lv base offs =
  if Ival.equal Ival.bottom offs
  then begin
      Relations_type.Model.bottom,
    (Some Zone.bottom),
    loc_bits_to_loc lv Location_Bits.bottom
    end
  else
    let result = match base with
    | Var host ->
        let base = Base.create_varinfo host in
        state, deps,
        loc_bits_to_loc lv (Location_Bits.inject base offs)
    | Mem x ->
        let state, deps, loc_lv =
          eval_expr_with_deps_state ~with_alarms deps state x
        in
        let loc_bits =
          Location_Bits.location_shift
            offs
            (loc_bytes_to_loc_bits loc_lv)
        in
        state, deps, loc_bits_to_loc lv loc_bits
    in
    CilE.set_syntactic_context (CilE.SyMem lv);
    result

and eval_expr ~with_alarms state e =
  snd (eval_expr_with_deps ~with_alarms None state e)

and get_influential_vars ~with_alarms state cond =
  (*  Format.printf "get_influential cond:%a@.state:%a@."
      !d_exp cond
      Relations_type.Model.pretty state; *)

  let rec get_vars acc cond =
    match cond.enode with
    | Lval (Var v, off as lv) ->
        let offset =
          try
            let _, _, offset =
              eval_offset ~reduce_valid_index:true ~with_alarms None
                v.vtype state off
            in
            offset
          with Offset_not_based_on_Null _ ->
            Ival.top
        in
        if Ival.cardinal_zero_or_one offset
        then
          (* no variable in offset can be influential *)
          let varid = Base.create_varinfo v in
          let loc =
            Locations.make_loc
              (Locations.Location_Bits.inject varid offset)
              (sizeof_lval lv)
          in
          let contents = Relations_type.Model.find state ~with_alarms loc in
          if Location_Bytes.cardinal_zero_or_one contents
          then (
              (*              Format.printf "cond:%a@.var contents:%a@.state:%a@."
                              !d_exp cond
                              Location_Bytes.pretty contents
                              Relations_type.Model.pretty state; *)

              acc (* it's not influential *)
            )
          else loc :: acc
        else
          (* a variable in offset can be influential *)
          get_vars_offset acc off
    | Lval (Mem e, off) ->
        get_vars_offset (get_vars acc e) off
    | BinOp(_,v1,v2,_) ->
        get_vars (get_vars acc v1) v2
    | UnOp(_,v1,_) ->
        get_vars acc v1
    | CastE (_typ,exp) ->
        get_vars acc exp
    | _ -> acc
  and get_vars_offset acc offset =
    match offset with
      NoOffset -> acc
    | Field (_,off) -> get_vars_offset acc off
    | Index (ind,off) -> get_vars (get_vars_offset acc off) ind
  in
  get_vars [] cond

and reduce_by_valid_expr ~with_alarms ~positive exp state =
  try
    ignore (with_alarms);
    let lv =
      match exp.enode with
        Lval lv -> lv
      | _ -> raise Cannot_find_lv
    in
    (* TODO: utiliser find_lv_plus pour traiter plus d'expressions *)
    let loc = lval_to_loc ~with_alarms:CilE.warn_none_mode state lv in
    if not (Locations.valid_cardinal_zero_or_one loc)
    then state
    else
      let value = Relations_type.Model.find
        ~with_alarms:CilE.warn_none_mode
        state
        loc
      in
      ( match value with
        Location_Bytes.Top _ ->
          (* we won't reduce anything anyway,
             and we may lose information if loc contains misaligned data *)

          raise Cannot_find_lv
      | _ -> () );
      let value_as_loc =
        make_loc
          (loc_bytes_to_loc_bits value)
          (sizeof_pointed (Cil.typeOfLval lv))
      in
      let reduced_value =
        loc_to_loc_without_size
          (if positive
            then valid_part value_as_loc
            else invalid_part value_as_loc )
      in
      if Location_Bytes.equal value reduced_value
      then state
      else begin
          if Location_Bytes.equal Location_Bytes.bottom reduced_value
          then Relations_type.Model.bottom
          else
            Relations_type.Model.reduce_binding
              state
              loc
              reduced_value
        end
  with Cannot_find_lv -> state

and eval_expr_with_deps ~with_alarms deps (state : Relations_type.Model.t) e =
  let _,deps,r = eval_expr_with_deps_state ~with_alarms deps state e in
  deps, r

and eval_BinOp ~with_alarms e deps state =
  match e.enode with
    BinOp (op, e1, e2, typ) ->
      let state, deps, ev1 =
        eval_expr_with_deps_state ~with_alarms deps state e1
      in
      if V.is_bottom ev1
      then Relations_type.Model.bottom, (Some Zone.bottom) ,V.bottom
      else
        let state, deps, ev2 =
          eval_expr_with_deps_state ~with_alarms deps state e2
        in
        if V.is_bottom ev2
        then Relations_type.Model.bottom, (Some Zone.bottom) ,V.bottom
        else begin
            let syntactic_context = CilE.SyBinOp (op,e1,e2) in
            CilE.set_syntactic_context syntactic_context;
            begin match unrollType (typeOf e1) with
            | TFloat _ ->
                let interpreted_expr =
                  (* refactor: shouldn't this be somewhere else? *)
                  try
                    let f1 =
                      try
                        let v1 = V.find_ival ev1 in
                        Ival.project_float v1
                      with V.Not_based_on_null
                      | Ival.Float_abstract.Nan_or_infinite ->
                          Value_parameters.warning ~current:true ~once:true
                            "float value must be finite: assert(TODO)";
                          Ival.Float_abstract.top
                    in
                    let f2 =
                      try
                        let v2 = V.find_ival ev2 in
                        Ival.project_float v2
                      with V.Not_based_on_null
                      | Ival.Float_abstract.Nan_or_infinite ->
                          Value_parameters.warning ~current:true ~once:true
                            "converting value to float: assert(TODO)";
                          Ival.Float_abstract.top
                    in
                    let binary_float_floats _name f =
                      try
                        let alarm, f = f (get_rounding_mode ()) f1 f2 in
                        if alarm then
                          CilE.warn_result_nan_infinite with_alarms ;
                        V.inject_ival (Ival.inject_float f)
                      with
                        Ival.Float_abstract.Nan_or_infinite ->
                          CilE.warn_result_nan_infinite with_alarms ;
                          V.top_float
                      | Ival.Float_abstract.Bottom ->
                          CilE.warn_result_nan_infinite with_alarms ;
                          V.bottom
                    in
                    begin match op with
                    | PlusA ->
                        binary_float_floats "+."
                          Ival.Float_abstract.add_float
                    | MinusA ->
                        binary_float_floats "-."
                          Ival.Float_abstract.sub_float
                    | Mult ->
                        binary_float_floats "*."
                          Ival.Float_abstract.mult_float
                    | Div ->
                        if Ival.Float_abstract.contains_zero f2
                        then
                          Value_parameters.warning ~once:true ~current:true
                            "float division: assert(TODO)";
                        binary_float_floats "/."
                          Ival.Float_abstract.div_float
                    | Eq ->
                        let contains_zero, contains_non_zero =
                          Ival.Float_abstract.equal_float_ieee f1 f2
                        in
                        V.interp_boolean ~contains_zero ~contains_non_zero
                    | Ne ->
                        let contains_non_zero, contains_zero =
                          Ival.Float_abstract.equal_float_ieee f1 f2
                        in
                        V.interp_boolean ~contains_zero ~contains_non_zero
                    | Lt ->
                        V.interp_boolean
                          ~contains_zero:(Ival.Float_abstract.maybe_le_ieee_float f2 f1)
                          ~contains_non_zero:(Ival.Float_abstract.maybe_lt_ieee_float f1 f2)
                    | Le ->
                        V.interp_boolean
                          ~contains_zero:(Ival.Float_abstract.maybe_lt_ieee_float f2 f1)
                          ~contains_non_zero:(Ival.Float_abstract.maybe_le_ieee_float f1 f2)
                    | Gt ->
                        V.interp_boolean
                          ~contains_zero:(Ival.Float_abstract.maybe_le_ieee_float f1 f2)
                          ~contains_non_zero:(Ival.Float_abstract.maybe_lt_ieee_float f2 f1)
                    | Ge ->
                        V.interp_boolean
                          ~contains_zero:(Ival.Float_abstract.maybe_lt_ieee_float f1 f2)
                          ~contains_non_zero:(Ival.Float_abstract.maybe_le_ieee_float f2 f1)
                    | _ -> raise V.Not_based_on_null
                    end
                  with V.Not_based_on_null | Ival.F.Nan_or_infinite ->
                    Value_parameters.warning ~once:true ~current:true "float operation on address: assert (TODO)";

                    V.join
                      (V.topify_arith_origin ev1)
                      (V.topify_arith_origin ev2)
                in
                state, deps, interpreted_expr
            | TInt _ | TPtr (_, _) | _ ->
                let interpreted_expr = begin match op with
                | PlusPI | IndexPI ->
                    V.add_untyped (osizeof_pointed typ) ev1 ev2
                | MinusPI ->
                    V.add_untyped (Int_Base.neg (osizeof_pointed typ)) ev1 ev2
                | PlusA ->
                    V.add_untyped (Int_Base.inject Int.one) ev1 ev2
                | MinusA | MinusPP ->
                    let minus_val = V.add_untyped Int_Base.minus_one ev1 ev2 in
                    if op = MinusA
                    then minus_val
                    else (* MinusPP *)
                      ( try
                          let size =
                            Int_Base.project (sizeof_pointed(Cil.typeOf e1))
                          in
                          let size = Int.div size (Int.of_int 8) in
                          if Int.equal size Int.one then
                            minus_val
                          else
                            let minus_val = Cvalue_type.V.find_ival minus_val in
                            Cvalue_type.V.inject_ival
                              (Ival.scale_div ~pos:true size minus_val)
                        with
                          Int_Base.Error_Top
                        | Cvalue_type.V.Not_based_on_null
                        | Not_found ->
                            V.join
                              (V.topify_arith_origin ev1)
                              (V.topify_arith_origin ev2))
                | Mod -> V.c_rem ~with_alarms ev1 ev2
                | Div -> V.div ~with_alarms ev1 ev2
                | Mult -> V.arithmetic_function ~with_alarms "*" Ival.mul ev1 ev2
                | LOr ->
                    assert false
                      (* This code makes a strict evaluation: V.interp_boolean
                         ~contains_zero: (V.contains_zero ev1 &&
                         V.contains_zero ev2) ~contains_non_zero:
                         (V.contains_non_zero ev1 || V.contains_non_zero
                         ev2)*)

                | LAnd ->
                    assert false
                      (* This code makes a strict evaluation:
                         V.interp_boolean ~contains_zero: (V.contains_zero
                         ev1 || V.contains_zero ev2) ~contains_non_zero:
                         (V.contains_non_zero ev1 && V.contains_non_zero
                         ev2)*)

                | BOr -> V.oper_on_values ~with_alarms "|" Int.logor ev1 ev2
                | BXor -> V.oper_on_values ~with_alarms "^" Int.logxor ev1 ev2
                | BAnd ->
                    ( try
                        let size = bitsSizeOf (typeOf e1)
                        in
                        V.bitwise_and ~size ev1 ev2
                      with SizeOfError _ ->
                        V.join
                          (V.topify_arith_origin ev1)
                          (V.topify_arith_origin ev2))

                | Eq | Ne | Ge | Le | Gt | Lt ->
                    let ev1, ev2 = check_comparable ~with_alarms ev1 ev2 in
                    let f = match op with
                    | Eq -> V.check_equal true
                    | Ne -> V.check_equal false
                    | Ge -> V.comparisons ">=" V.do_ge
                    | Le -> V.comparisons "<=" V.do_le
                    | Gt -> V.comparisons ">" V.do_gt
                    | Lt -> V.comparisons "<" V.do_lt
                    | _ -> assert false
                    in
                    f ev1 ev2
                | Shiftrt ->
                    begin try
                        let signed = is_signed_int_enum_pointer typ in
                        V.shift_right ~with_alarms ~size:(bitsSizeOf typ) ~signed ev1 ev2
                      with SizeOfError _ ->
                        (match with_alarms.CilE.imprecision_tracing with
                        | CilE.Aignore -> ()
                        | CilE.Acall f -> f ()
                        | CilE.Alog -> Value_parameters.result "shifting value of unknown size");
                        V.top  (* TODO: topify ... *)
                    end
                | Shiftlt ->
                    begin try
                        V.shift_left ~with_alarms ~size:(bitsSizeOf typ) ev1 ev2
                      with SizeOfError _ ->
                        (match with_alarms.CilE.imprecision_tracing with
                        | CilE.Aignore -> ()
                        | CilE.Acall f -> f ()
                        | CilE.Alog -> Value_parameters.result "shifting value of unknown size");
                        V.top (* TODO: topify ... *)
                    end
                  end
                in
                (* Warn if overflow in a signed int binop *)
                let interpreted_expr =
                  match op with
                    Shiftlt|Mult|MinusPP|MinusPI|IndexPI|
                        PlusPI|PlusA|Div|Mod|MinusA ->
                          handle_signed_overflow
                            ~with_alarms
                            syntactic_context
                            typ
                            e
                            interpreted_expr
                  | _ -> interpreted_expr
                in
                state, deps, interpreted_expr
            end
          end
  | _ -> assert false

and eval_expr_with_deps_state
    ~with_alarms deps (state : Relations_type.Model.t) e =
  let state, deps, expr =
    let orig_expr = Cil.stripInfo e in
    match orig_expr.enode with
    | Info _ -> assert false
    | Const v ->
        let r =
          begin match v with
          | CInt64 (i,k,_s) ->
              V.inject_int (
                  if isSigned k then Int.of_int64 i
                  else (* For 64 bits type we need to reinterpret the sign *)
                    let s = Printf.sprintf "%Lu" i in
                    Int.of_string s)
          | CChr c ->
              (match charConstToInt c with
              | CInt64 (i,_,_) -> V.inject_int (Int.of_int64 i)
              | _ -> assert false)
          | CReal (f, _fsize, _) ->
              Value_parameters.result ~once:true
                "float support is experimental";
              let f = Ival.F.of_float f in
              let overflow, af =
                try
                  Ival.Float_abstract.inject_r f f
                with Ival.Float_abstract.Bottom -> assert false
              in
              assert (not overflow);
              V.inject_ival (Ival.inject_float af)
          | CWStr _ ->
              Value_parameters.result "approximation because of a wide string";
              (* TODO *) V.top_int
          | CStr s ->
              V.inject (Base.create_string s) Ival.zero
          | CEnum {eival = e} ->
              let _,_, r =
                eval_expr_with_deps_state ~with_alarms deps state e
              in
              r
          end
        in
        state, deps, r
    | BinOp _  ->
        eval_BinOp ~with_alarms orig_expr deps state
    | Lval lv ->
        eval_lval ~with_alarms deps state lv
    | AddrOf v | StartOf v ->
        let state, deps, r =
          lval_to_loc_with_offset_deps_only_option ~with_alarms ?deps state v
        in
        state, deps, loc_to_loc_without_size r

    | CastE (typ, e) ->
        let deps, evaled_expr =
          eval_expr_with_deps ~with_alarms deps state e
        in
        let src_typ = unrollType (typeOf e) in
        let dest_type = unrollType typ in
        state, deps, do_promotion ~with_alarms ~dest_type ~src_typ evaled_expr

    | SizeOf typ ->
        let r =
          try V.inject_ival
            (Ival.inject_singleton ((Int.of_int ((bitsSizeOf typ) / 8))))
          with SizeOfError _ ->
            error "cannot interpret sizeof(incomplete type)";
            V.top_int
        in
        state, deps, r
    | SizeOfE e ->
        let typ = typeOf e in
        let r =
          try V.inject_ival
            (Ival.inject_singleton ((Int.of_int ((bitsSizeOf typ) / 8))))
          with SizeOfError _ ->
            error "cannot interpret sizeof(incomplete type)";
            V.top_int
        in
        state, deps, r

    | UnOp (LNot, e, _) ->
        (* TODO:  on float, LNot is equivalent to == 0.0 *)
        let deps, expr = eval_expr_with_deps ~with_alarms deps state e in
        CilE.set_syntactic_context (CilE.SyBinOp (EqCil.zero, e));
        let _, expr =
          check_comparable ~with_alarms V.singleton_zero expr
        in
        CilE.set_syntactic_context (CilE.SyUnOp e);
        let t1 = typeOf e in
        if isIntegralType t1 || isPointerType t1
        then
          state, deps, V.interp_boolean
            ~contains_zero:(V.contains_non_zero expr)
            ~contains_non_zero:(V.contains_zero expr)
        else state, deps, V.zero_or_one

    | UnOp (Neg, e, t) ->
        let t = unrollType t in
        let deps, expr = eval_expr_with_deps ~with_alarms deps state e in
        let syntactic_context = CilE.SyUnOp orig_expr in
        CilE.set_syntactic_context syntactic_context;
        ( match t with TFloat _ ->
          let result =
            try
              let v = V.find_ival expr in
              let f =
                Ival.project_float v
              in
              V.inject_ival
                (Ival.inject_float (Ival.Float_abstract.neg_float f))
            with
              V.Not_based_on_null ->
                begin match with_alarms.CilE.others with
                  CilE.Aignore -> ()
                | CilE.Acall f -> f()
                | CilE.Alog ->
                    Value_parameters.warning ~once:true ~current:true
                      "converting address to float: assert(TODO)"
                end;
                V.topify_arith_origin expr
            | Ival.Float_abstract.Nan_or_infinite ->
                begin match with_alarms.CilE.others with
                  CilE.Aignore -> ()
                | CilE.Acall f -> f()
                | CilE.Alog ->
                    Value_parameters.warning ~once:true ~current:true
                      "converting value to float: assert (TODO)"
                end;
                V.top_float
          in
          state, deps, result
        | _ ->
            let result =
              try
                let v = V.find_ival expr in
                V.inject_ival (Ival.neg v)
              with V.Not_based_on_null -> V.topify_arith_origin expr
            in
            let result =
              handle_signed_overflow ~with_alarms
                syntactic_context t orig_expr result
            in
            state, deps, result)

    | UnOp (BNot, e, _) ->
        let deps, expr = eval_expr_with_deps ~with_alarms deps state e in
        CilE.set_syntactic_context (CilE.SyUnOp e);
        let result =
          try
            let v = V.find_ival expr in
            V.inject_ival
              (Ival.apply_set_unary "~" Int.lognot v)
          with V.Not_based_on_null -> V.topify_arith_origin expr
        in
        state, deps, result
    | AlignOfE _|AlignOf _|SizeOfStr _
        ->
        Value_parameters.result
          "C construct alignof or sizeof string not precisely handled";
          state, deps, V.top_int
  in
  let r =
    if hasAttribute "volatile" (typeAttrs (typeOf e))
      && not (Cvalue_type.V.equal Cvalue_type.V.bottom expr)
    then V.top_int
    else
      expr
  in
  let state,r_ptr = PtrRelational.eval_expr ~with_alarms state e in
  let r = Cvalue_type.V.narrow r_ptr r in
  let r = do_cast ~with_alarms (typeOf e) r in
  state, deps, r

and eval_expr_with_deps_state_subdiv ~with_alarms deps
    (state : Relations_type.Model.t) e =
  let ((state_without_subdiv, deps_without_subdiv, result_without_subdiv) as result) =
    eval_expr_with_deps_state  ~with_alarms deps
      (state : Relations_type.Model.t) e
  in
  let subdivnb = Value_parameters.Subdivide_float_in_expr.get() in
  if subdivnb=0
  then result
  else if not (Locations.Location_Bytes.is_included result_without_subdiv Locations.Location_Bytes.top_int)
  then begin
      Value_parameters.debug ~level:2
        "subdivfloatvar: expression has an address result";
      result
    end
  else
    let compare_min, compare_max =
      if Locations.Location_Bytes.is_included result_without_subdiv Locations.Location_Bytes.top_float
      then begin
          Value_parameters.debug ~level:2
            "subdivfloatvar: optimizing floating-point expression %a=%a"
            !d_exp e
            Locations.Location_Bytes.pretty result_without_subdiv;
          Cvalue_type.V.compare_min_float, Cvalue_type.V.compare_max_float
        end
      else begin
          Value_parameters.debug ~level:2
            "subdivfloatvar: optimizing integer expression %a=%a"
            !d_exp e
            Locations.Location_Bytes.pretty result_without_subdiv;
          Cvalue_type.V.compare_min_int, Cvalue_type.V.compare_max_int
        end
    in
    let vars =
      get_influential_vars ~with_alarms:CilE.warn_none_mode state e
    in
    Value_parameters.debug ~level:2 "subdivfloatvar: variable list=%a"
      (Pretty_utils.pp_list Locations.pretty)
      vars;
    let rec try_sub vars =
      match vars with
        [] | [ _ ] ->
          result
      | v :: tail ->
          try
            if not (List.exists (fun x -> Locations.loc_equal v x) tail)
            then raise Too_linear;
            let v_value =
              Relations_type.Model.find
                ~with_alarms:CilE.warn_none_mode
                state
                v
            in
            Value_parameters.debug ~level:2
              "subdivfloatvar: considering optimizing variable %a (value %a)"
              Locations.pretty v Cvalue_type.V.pretty v_value;
            if not (Locations.Location_Bytes.is_included v_value Locations.Location_Bytes.top_float)
            then raise Too_linear;

            let working_list = ref [ (v_value, result_without_subdiv) ] in
            let had_bottom = ref false in
            let subdiv_for_bound better_bound =
              let rec insert_subvalue_in_list (_, exp_value as p) l =
                match l with
                  [] -> [p]
                | (_, exp_value1 as p1) :: tail ->
                    if better_bound exp_value1 exp_value >= 0
                    then p :: l
                    else p1 :: (insert_subvalue_in_list p tail)
              in
              let exp_subvalue subvalue l =
                let substate =
                  (* FIXME: should be relation-aware primitive *)
                  Relations_type.Model.add_binding
                    ~with_alarms:CilE.warn_none_mode
                    ~exact:true
                    state
                    v
                    subvalue
                in
                let subexpr = eval_expr ~with_alarms substate e in
                Value_parameters.debug ~level:2
                  "subdivfloatvar: computed var=%a expr=%a"
                  V.pretty subvalue
                  V.pretty subexpr;
                if Cvalue_type.V.is_bottom subexpr
                then begin
                    had_bottom := true;
                    l
                  end
                else
                  insert_subvalue_in_list (subvalue, subexpr) l
              in
              let subdiv l =
                match l with
                  [] ->
                    Value_parameters.debug
                      "subdivfloatvar: all reduced to bottom!!";
                    raise Ival.Can_not_subdiv
                | (value, _exp_value) :: tail ->
                    let (subvalue1, subvalue2) =
                      Cvalue_type.V.subdiv_float_interval value
                    in
                    let s = exp_subvalue subvalue1 tail
                    in
                    exp_subvalue subvalue2 s
              in
              try
                for i = 1 to subdivnb do
                  working_list := subdiv !working_list;
                done
              with Ival.Can_not_subdiv -> ()
            in
            subdiv_for_bound compare_min ;
            (* sort working_list in decreasing order
               on the upper bounds of exp_value *)

            let comp_exp_value (_value1,exp_value1) (_value2,exp_value2) =
              compare_max exp_value1 exp_value2
            in
            working_list := List.sort comp_exp_value !working_list ;
            if Value_parameters.debug_atleast 2 then
              List.iter
                (function (x, e) ->
                  Value_parameters.debug
                    "subdivfloatvar: elements of list max %a %a"
                    V.pretty x V.pretty e)
                !working_list;
            subdiv_for_bound compare_max ;
            let working_list = !working_list in
            if Value_parameters.debug_atleast 2 then
              List.iter
                (function (x, e) ->
                  Value_parameters.debug
                    "subdivfloatvar: elements of final list %a %a"
                    V.pretty x V.pretty e)
                working_list;
            let reduced_state, optimized_exp_value =
              if !had_bottom
              then
                let reduced_var, optimized_exp_value =
                  List.fold_left
                    (fun (accv,acce) (value, exp_value)  ->
                      Cvalue_type.V.join value accv,
                      Cvalue_type.V.join exp_value acce)
                    (Cvalue_type.V.bottom,
                    Cvalue_type.V.bottom)
                    working_list
                in
                (* FIXME: should be relation-aware primitive *)
                Relations_type.Model.add_binding
                  ~with_alarms:CilE.warn_none_mode
                  ~exact:true
                  state
                  v
                  reduced_var,
              optimized_exp_value
              else
                state_without_subdiv,
              List.fold_left
                (fun acc (_value, exp_value)  ->
                  Cvalue_type.V.join exp_value acc)
                Cvalue_type.V.bottom
                working_list
            in
            reduced_state, deps_without_subdiv, optimized_exp_value
          with Not_less_than | Too_linear ->
            try_sub tail
    in
    try_sub vars

and eval_lval_using_main_memory ~with_alarms
    deps (state:Relations_type.Model.t) lv
    =
  let state,deps,loc =
    lval_to_loc_deps_option ~with_alarms ?deps state lv
      ~reduce_valid_index:(Parameters.SafeArrays.get ())
  in
  CilE.set_syntactic_context (CilE.SyMem lv);
  let result = Relations_type.Model.find ~with_alarms state loc in
  (* TODO: move into Model.find *)
  let valid_loc = Locations.valid_part loc in
  let state =
    if Location_Bits.equal loc.Locations.loc valid_loc.Locations.loc
    then state
    else begin
        match lv with
          (*          Mem (Lval ((_,_) as lv_mem)),NoOffset ->
                  let loc_mem =
                  lval_to_loc ~with_alarms:warn_none_mode state lv_mem
                  in
                  if Location_Bits.cardinal_zero_or_one loc_mem.Locations.loc
                  then Relations_type.Model.reduce_binding
                  state loc_mem
                  (loc_bits_to_loc_bytes valid_loc.loc)
                  else state *)

          Mem (exp_mem),NoOffset ->
            let lv_mem_plus_list =
              find_lv_plus ~with_alarms:CilE.warn_none_mode state exp_mem
            in
            let treat_lv_mem_plus (lv_mem, plus) state =
              let loc_mem =
                lval_to_loc ~with_alarms:CilE.warn_none_mode state lv_mem
              in
              if Location_Bits.cardinal_zero_or_one loc_mem.Locations.loc
              then
                let new_val =
                  Location_Bytes.location_shift
                    (Ival.neg plus)
                    (loc_bits_to_loc_bytes valid_loc.loc)
                in
                Relations_type.Model.reduce_binding
                  state loc_mem new_val
              else state
            in
            List.fold_right treat_lv_mem_plus lv_mem_plus_list state
        | _ -> state
      end
  in
  (match with_alarms.CilE.imprecision_tracing with
  | CilE.Aignore -> ()
  | CilE.Acall f -> f ()
  | CilE.Alog -> warn_lval_read lv loc result);
  let new_deps =
    match deps with
    | None -> None
    | Some deps -> Some (Zone.join deps (valid_enumerate_bits loc))
  in
  state, new_deps, result

and eval_lval ~with_alarms deps state (base,offset as lv) =
  let state, deps, result_from_main_memory =
    eval_lval_using_main_memory ~with_alarms deps state lv
  in
  let find_loc_mem sub_lv offs =
    try
      let loc = lval_to_loc ~with_alarms state sub_lv in
      let size = sizeof_lval lv in
      CilE.set_syntactic_context (CilE.SyMem lv);
      Relations_type.Model.find_mem loc size offs state
    with Relations_type.Use_Main_Memory ->
      result_from_main_memory
  in
  let result = match base with
  | Mem({enode = Lval sub_lv} as e) when UseRelations.get () ->
      let typ = typeOf e in
      begin try
          let _, _, offs =
            eval_offset ~reduce_valid_index:(Parameters.SafeArrays.get ())
              ~with_alarms None typ state offset
          in
          find_loc_mem sub_lv offs
        with
          Offset_not_based_on_Null _ ->
            result_from_main_memory
      end
  | Mem({enode = BinOp((PlusPI|IndexPI|MinusPI as op),
                      {enode = Lval sub_lv} ,
                      e2,_)}
           as e)
      when UseRelations.get () ->
      begin
        let e2 = eval_expr ~with_alarms state e2 in
        let typ = typeOf e in
        try
          let ival = Cvalue_type.V.find_ival e2 in
          let ival = if op=MinusPI then Ival.neg ival else ival in
          let _, _, offs =
            eval_offset ~reduce_valid_index:(Parameters.SafeArrays.get ())
              ~with_alarms None typ state offset in
          let offs = (* convert to bits *)
            Ival.add
              (Ival.scale
                  (Int_Base.project (sizeof_pointed typ))
                  ival)
              offs
          in
          let result = find_loc_mem sub_lv offs in
          result
        with
        | Offset_not_based_on_Null _
        | Int_Base.Error_Top
        | Cvalue_type.V.Not_based_on_null -> result_from_main_memory
      end
  | _e ->
      result_from_main_memory
  in
  let result_inter = Cvalue_type.V.narrow result_from_main_memory result in
  state, deps, result_inter

and eval_offset ~reduce_valid_index ~with_alarms deps typ state offset =
  match offset with
  | NoOffset ->
      state, deps, Ival.singleton_zero
  | Index (exp,remaining) ->
      let typ_pointed,array_size = match (unrollType typ) with
      | TArray (t,size,_,_) -> t, size
      | TPtr(t,_) ->
          (match unrollType t with
          | TArray (t,size,_,_) -> t,size (* pointer to start of an array *)
          | _ ->
              error "Got type '%a'" !Ast_printer.d_type t;
              assert false)
      | t ->
          error "Got type '%a'" !Ast_printer.d_type t;
          assert false
      in
      let state, deps, current =
        eval_expr_with_deps_state ~with_alarms deps state exp
      in
      if V.is_bottom current
      then Relations_type.Model.bottom, (Some Zone.bottom), Ival.bottom
      else
        let state, offset =
          try
            let v = V.find_ival current in
            let state, v =
              if reduce_valid_index then
                try
                  let array_siz = lenOfArray64 array_size in
                  let new_v =
                    Ival.narrow (Ival.inject_range
                                    (Some Int.zero)
                                    (Some (Int.of_int64 (Int64.pred array_siz)))) v
                  in
                  let new_state =
                    if Ival.equal new_v v
                    then state
                    else begin
                        begin
                          match with_alarms.CilE.others with
                          | CilE.Aignore -> ()
                          | CilE.Acall f -> f ()
                          | CilE.Alog ->
                              CilE.set_syntactic_context
                                (CilE.SyBinOp
                                    (IndexPI,
                                    exp,
                                    Cilutil.out_some array_size));
                              CilE.warn_index  with_alarms "accessing"
                        end;
                        state (* TODO : if the index is a variable, reduce *)
                      end
                  in
                  new_state, new_v
                with LenOfArray -> state, v
              else state, v
            in
            state, v
          with V.Not_based_on_null ->
            let deps, offset =
              topify_offset
                ~with_alarms
                deps
                state
                (Cvalue_type.V.topify_arith_origin current)
                remaining
            in
            raise (Offset_not_based_on_Null (deps,offset))
        in
        let state, deps, r =
          eval_offset ~reduce_valid_index ~with_alarms
            deps typ_pointed state remaining
        in
        let offset = Ival.scale_int64base (sizeof typ_pointed) offset in
        state, deps, Ival.add offset r
  | Field (fi,remaining) ->
      let current,_ = bitsOffset typ (Field(fi,NoOffset)) in
      let state, deps, r =
        eval_offset ~reduce_valid_index ~with_alarms
          deps
          fi.ftype
          state
          remaining
      in
      state, deps, Ival.add (Ival.of_int current) r
and topify_offset ~with_alarms deps state acc offset =
  match offset with
  | NoOffset -> deps,acc
  | Field (_fi,remaining) -> topify_offset ~with_alarms deps state acc remaining
  | Index (exp,remaining) ->
      let deps, loc_index = eval_expr_with_deps ~with_alarms deps state exp in
      let acc = Location_Bytes.join
        (Cvalue_type.V.topify_arith_origin loc_index)
        acc
      in
      topify_offset ~with_alarms deps state acc remaining