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

and exp_lval_to_loc state exp =
  let lv =
    match exp.enode with
      | Lval lv -> lv
      | _ -> raise Cannot_find_lv
  in
    (* TODO: utiliser find_lv_plus pour traiter plus d'expressions *)
  lv, lval_to_loc ~with_alarms:CilE.warn_none_mode state lv

and lval_to_loc_deps_option
    ~with_alarms ~deps state ~reduce_valid_index (base,offset as lv)
 =
  if not (Cvalue.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 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:Cvalue.Model.t) (_base, _offset as v)
    =
  lval_to_loc_deps_option
    ~with_alarms ~deps (state:Cvalue.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 not typ_ge_typeofe then raise exn;
        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.project_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
      with Neither_Int_Nor_Enum_Nor_Pointer
        -> raise exn)

and find_lv ~with_alarms (state:Cvalue.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 (FDouble,_), 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.project_ival ev1 in
                          find_lv_plus_rec e2 (Ival.add current_offs ival1)
                        with V.Not_based_on_null -> ());
                      ( try
                          let ival2 = V.project_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.project_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.is_bottom offs
  then begin
      Cvalue.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.find 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
      Cvalue.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 =
            Cvalue.Model.find ~conflate_bottom:true
              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
                 Cvalue.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 ~positive ~for_writing exp state =
  try
    let lv, loc = exp_lval_to_loc state exp in
    if not (Locations.valid_cardinal_zero_or_one ~for_writing:false loc)
    then state
    else reduce_by_valid_loc ~positive ~for_writing loc (typeOfLval lv) state
  with Cannot_find_lv -> state

and reduce_by_valid_loc ~positive ~for_writing loc typ state =
  try
    let value = Cvalue.Model.find ~with_alarms:CilE.warn_none_mode
      ~conflate_bottom:true state loc
    in
    if Cvalue.V.is_imprecise value then
      (* 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 typ)
    in
    let reduced_value =
      loc_to_loc_without_size
        (if positive
         then valid_part ~for_writing 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 Cvalue.Model.bottom
      else
        Cvalue.Model.reduce_binding ~with_alarms:CilE.warn_none_mode
          state loc reduced_value
    end
  with Cannot_find_lv -> state

and reduce_by_initialized_loc ~with_alarms:_ ~positive (typ, loc) state =
  try
    let locbi = loc_bytes_to_loc_bits loc in
    let size = match unrollType typ with
      | TPtr (t, _) -> bitsSizeOf t
      | _ -> assert false
    in
    let loc = make_loc locbi (Int_Base.inject (Int.of_int size)) in
    if not (Locations.valid_cardinal_zero_or_one ~for_writing:false loc)
    then state
    else
      let value = Cvalue.Model.find_unspecified
        ~with_alarms:CilE.warn_none_mode state loc
      in
      (match value with
      | Cvalue.V_Or_Uninitialized.C_uninit_esc (Location_Bytes.Top _)
      | Cvalue.V_Or_Uninitialized.C_uninit_noesc (Location_Bytes.Top _)
      | Cvalue.V_Or_Uninitialized.C_init_esc (Location_Bytes.Top _)
      | Cvalue.V_Or_Uninitialized.C_init_noesc (Location_Bytes.Top _) ->
          (* we won't reduce anything anyway,
             and we may lose information if loc contains misaligned data *)

          raise Cannot_find_lv
      | _ -> () );
      let reduced_value =
        Cvalue.V_Or_Uninitialized.change_initialized positive value
      in
      if Cvalue.V_Or_Uninitialized.equal value reduced_value
      then state
      else begin
          if Cvalue.V_Or_Uninitialized.equal
            Cvalue.V_Or_Uninitialized.bottom reduced_value
          then Cvalue.Model.bottom
          else
            Cvalue.Model.add_binding_unspecified
              state
              loc
              reduced_value
        end
  with Cannot_find_lv -> state

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 Cvalue.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 Cvalue.Model.bottom, (Some Zone.bottom) ,V.bottom
        else begin
            begin match unrollType (typeOf e1) with
            | TFloat _ ->
                CilE.set_syntactic_context (CilE.SyUnOp e);
                let r = eval_binop_float ~with_alarms ev1 op ev2 in
                state, deps, r
            | TInt _ | TPtr (_, _) | _ as te1 ->
               CilE.set_syntactic_context (CilE.SyBinOp (op, e1, e2));
               let v = eval_binop_int
                 ~with_alarms ~typ ~te1 ev1 op ev2 in
               (* Warn if overflow in a signed int binop *)
               let v = match op with
                 | Shiftlt | Mult | MinusPP | MinusPI | IndexPI | PlusPI
                 | PlusA | Div | Mod | MinusA ->
                     handle_signed_overflow ~with_alarms typ v
                 | _ -> v
              in
               state, deps, v
            end
          end
  | _ -> assert false

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

and eval_expr_with_deps_state ~with_alarms deps state e =
  let state, deps, r =
    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 i
          | CChr c ->
              (match charConstToInt c with
              | CInt64 (i,_,_) -> V.inject_int 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
                  let o, af = Ival.Float_abstract.inject_r f f in
                  o, V.inject_ival (Ival.inject_float af)
                with Ival.Float_abstract.Bottom ->
                  Value_parameters.result ~current:true
                    "Floating-point literal (or constant expression) is not finite. This path is assumed to be dead.";
                  trueV.bottom
              in
              if overflow
              then Value_parameters.result "overflow in constant: assert(Ook);";
              af
          | CWStr _ | CStr _ ->
              V.inject (Base.create_string orig_expr) Ival.zero
          | CEnum {eival = e} ->
              eval_expr ~with_alarms state e
          end
        in
        state, deps, r
    | BinOp _  ->
        eval_BinOp ~with_alarms orig_expr deps state
    | Lval lv ->
        eval_lval ~conflate_bottom:true ~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
        let r = do_promotion ~with_alarms ~dest_type ~src_typ evaled_expr e in
        state, deps, r

    | SizeOf _ | SizeOfE _ | SizeOfStr _ | AlignOf _ | AlignOfE _ ->
        let e = Cil.constFold true orig_expr in
        let r = match e.enode with
          | Const (CInt64 (v, _, _)) -> Cvalue.V.inject_int v
          | _ ->
            Value_parameters.error ~current:true
              "cannot interpret sizeof or alignof (incomplete type)";
            V.top_int
        in
        state, deps, r

    | UnOp (op, e, _t_res) ->
        let t = unrollType (typeOf e) in
        let deps, expr = eval_expr_with_deps ~with_alarms deps state e in
        let syntactic_context = match op with
          | Neg -> CilE.SyUnOp orig_expr (* Can overflow *)
          | BNot -> CilE.SyUnOp orig_expr (* does in fact never raise an alarm*)
          | LNot -> CilE.SyBinOp (EqCil.zero ~loc:e.eloc, e)
              (* Can raise a pointer comparison. CilE needs a binop there *)
        in
        CilE.set_syntactic_context syntactic_context;
        let result = eval_unop ~with_alarms expr t op in
        state, deps, result
  in
  let r =
    if hasAttribute "volatile" (typeAttrs (typeOf e))
      && not (Cvalue.V.is_bottom r)
    then V.top_int
    else
      r
  in
  let typ = typeOf e in
  let rr = do_cast ~with_alarms typ r in
(*  ( match typ with
    TInt _ when not (V.equal r rr || V.is_topint r) ->
        warning_once_current
          "downcast %a -> %a@." V.pretty r V.pretty rr
  | _ -> ()); *)

  state, deps, rr

and eval_unop ~with_alarms expr t op =
  match op with
    | Neg ->
        let t = unrollType t in
        (match t with TFloat _ ->
           (try
              let v = V.project_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 _ ->
                    warning_once_current
                      "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 _ ->
                    warning_once_current
                      "converting value to float: assert (TODO)"
                end;
                V.top_float
           )
        | _ ->
            let result =
              try
                let v = V.project_ival expr in
                V.inject_ival (Ival.neg v)
              with V.Not_based_on_null -> V.topify_arith_origin expr
            in
            handle_signed_overflow ~with_alarms t result
        )

    | BNot ->
        (try
           let v = V.project_ival expr in
           V.inject_ival
             (Ival.apply_set_unary "~" Int.lognot v)
         with V.Not_based_on_null -> V.topify_arith_origin expr)

    | LNot ->
        (* TODO:  on float, LNot is equivalent to == 0.0 *)
        let warn, _, expr = check_comparable Eq V.singleton_zero expr in
        if warn then CilE.warn_pointer_comparison with_alarms;
        if (warn &&
              Value_parameters.UndefinedPointerComparisonPropagateAll.get ())
          || not (isIntegralType t || isPointerType t)
        then
          V.zero_or_one
        else
          V.interp_boolean
            ~contains_zero:(V.contains_non_zero expr)
            ~contains_non_zero:(V.is_included V.singleton_zero expr)

and eval_expr_with_deps_state_subdiv ~with_alarms deps state e =
  let (state_without_subdiv, deps_without_subdiv, result_without_subdiv as r) =
    eval_expr_with_deps_state ~with_alarms deps state e
  in
  let subdivnb = Value_parameters.Subdivide_float_in_expr.get() in
  if subdivnb=0
  then r
  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";
      r
    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.V.compare_min_float, Cvalue.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.V.compare_min_int, Cvalue.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
      | [] | [ _ ] -> r
      | v :: tail ->
          try
            if not (List.exists (fun x -> Locations.loc_equal v x) tail)
            then raise Too_linear;
            let v_value =
              Cvalue.Model.find
                ~conflate_bottom:true
                ~with_alarms:CilE.warn_none_mode
                state
                v
            in
    (*        Value_parameters.result ~current:true
              "subdivfloatvar: considering optimizing variable %a (value %a)"
              Locations.pretty v Cvalue.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 *)
                  Cvalue.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.result ~current:true
                  "subdivfloatvar: computed var=%a expr=%a"
                  V.pretty subvalue
                  V.pretty subexpr; *)

                if Cvalue.V.is_bottom subexpr
                then begin
                    had_bottom := true;
                    l
                  end
                else
                  insert_subvalue_in_list (subvalue, subexpr) l
              in
              let size =
                if Value_parameters.AllRoundingModes.get ()
                then 0
                else Int.to_int (Int_Base.project v.Locations.size)
              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.V.subdiv_float_interval ~size 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 ;
            (* Now 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.V.join value accv,
                      Cvalue.V.join exp_value acce)
                    (Cvalue.V.bottom,
                    Cvalue.V.bottom)
                    working_list
                in
                (* FIXME: should be relation-aware primitive *)
                Cvalue.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.V.join exp_value acc)
                Cvalue.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

(* TODO. Cf also [Eval_stmts.do_assign_abstract_value_to_loc]. Should we check
   that the value is an int, and topify otherwise ? *)

and cast_lval_bitfield lv size v =
  let signed = signof_typeof_lval lv in
  Cvalue.V.cast ~with_alarms:CilE.warn_none_mode ~size ~signed v

and cast_lval_when_bitfield lv ?(sizelv=bitfield_size_lv lv) ?(sizebf=(bitfield_size_bf lv)) v =
  match sizebf with
    | Int_Base.Value size when is_bitfield lv ~sizelv ~sizebf () ->
        cast_lval_bitfield lv size v
    | _ -> v

and eval_lval_using_main_memory ~conflate_bottom ~with_alarms deps state lv =
  let state,deps,loc =
    lval_to_loc_deps_option ~with_alarms ?deps state lv
      ~reduce_valid_index:(Kernel.SafeArrays.get ())
  in
  CilE.set_syntactic_context (CilE.SyMem lv);
  let result =
    Cvalue.Model.find ~conflate_bottom ~with_alarms state loc
  in
(*  Format.printf "lval %a before %a@."
    !d_lval lv
    Cvalue.V.pretty result; *)

  let result = cast_lval_when_bitfield lv ~sizebf:loc.Locations.size result in
(*  Format.printf "lval %a after %a@."
    !d_lval lv
    Cvalue.V.pretty result; *)

  (* TODO: move into Model.find *)
  let valid_loc = Locations.valid_part ~for_writing:false loc in
  let state =
    if Location_Bits.equal loc.Locations.loc valid_loc.Locations.loc
    then state
    else begin
        match lv with
          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.is_relationable loc_mem.Locations.loc
              then
                let new_val =
                  Location_Bytes.location_shift
                    (Ival.neg plus)
                    (loc_bits_to_loc_bytes valid_loc.loc)
                in
                Cvalue.Model.reduce_binding
                  ~with_alarms:CilE.warn_none_mode
                  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 ~for_writing:false loc))
  in
  state, new_deps, result

and eval_lval ~conflate_bottom ~with_alarms deps state lv =
  let state, deps, result =
    eval_lval_using_main_memory ~conflate_bottom ~with_alarms deps state lv
  in
  let result_conv =
    (*    match unrollType (Cil.typeOfLval lv) with
          TFloat (FDouble|FFloat as kind, _) ->
          let f, r = Cvalue.V.force_float kind result in
          if f then Format.printf "TODO: assert@.";
          r
          | _ -> *)
 result
  in
  state, deps, result_conv

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 *)
          | _ ->
              Value_parameters.error ~current:true
                "Got type '%a'" !Ast_printer.d_type t;
              assert false)
      | t ->
          Value_parameters.error ~current:true
            "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 Cvalue.Model.bottom, (Some Zone.bottom), Ival.bottom
      else
        let state, offset =
          try
            let v = V.project_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 (My_bigint.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"
                                (Pretty_utils.sfprintf "%a" V.pretty current);
                        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.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.V.topify_arith_origin loc_index)
        acc
      in
      topify_offset ~with_alarms deps state acc remaining