diff --git a/src/plugins/wp/CfgCompiler.ml b/src/plugins/wp/CfgCompiler.ml
index e3a52d9d39cce44ced83687bad1522e15e4bbfd1..b5fee56975b07a8595d49f658f9aaad70b0aab60 100644
--- a/src/plugins/wp/CfgCompiler.ml
+++ b/src/plugins/wp/CfgCompiler.ml
@@ -20,7 +20,8 @@
 (*                                                                        *)
 (**************************************************************************)
 
-open Sigs
+open Memory
+open Sigma
 open Cil_types
 open Lang
 
@@ -33,1426 +34,1333 @@ type mode = [
   | `Bool_Forward
 ]
 
-module type Cfg =
-sig
-
-  module S : Sigma
-
-  module Node : sig
-    type t
-    module Map : Qed.Idxmap.S with type key = t
-    module Set : Qed.Idxset.S with type elt = t
-    module Hashtbl : Hashtbl.S with type key = t
-    val pp: Format.formatter -> t -> unit
-    val create: unit -> t
-    val equal: t -> t -> bool
-  end
-
-  type node = Node.t
-
-  val node : unit -> node
+module Node : sig
+  type t
+  module Map : Qed.Idxmap.S with type key = t
+  module Set : Qed.Idxset.S with type elt = t
+  module Hashtbl : FCHashtbl.S with type key = t
+  val tag: t -> int
+  val compare: t -> t -> int
+  val equal: t -> t -> bool
+  val pp: Format.formatter -> t -> unit
+  val create: unit -> t
+  val node_internal: unit -> t
+end
+= struct
+  type t = int
+  module I = struct type t = int let id x = x end
+  module Map = Qed.Idxmap.Make(I)
+  module Set = Qed.Idxset.Make(I)
+  module Hashtbl = Datatype.Int.Hashtbl
+  let tag = I.id
+  let compare = Datatype.Int.compare
+  let equal = Datatype.Int.equal
+  let pp fmt n =
+    if n>=0 then Format.pp_print_int fmt n
+    else Format.fprintf fmt "int%i" (-n)
+
+  let node_compter = ref (-1)
+
+  let create () =
+    incr node_compter;
+    !node_compter
+
+  let node_internal_compter = ref 0
+
+  let node_internal () =
+    decr node_internal_compter;
+    !node_internal_compter
 
-  module C :
-  sig
-    type t
-    val equal : t -> t -> bool
-    val create : S.t -> F.pred -> t
-    val get : t -> F.pred
-    val reads : t -> S.domain
-    val relocate : S.t -> t -> t
-  end
+end
 
-  module P :
-  sig
-    type t
-    val pretty : Format.formatter -> t -> unit
-    val create : S.t Node.Map.t -> F.pred -> t
-    val get: t -> F.pred
-    val reads : t -> S.domain Node.Map.t
-    val nodes : t -> Node.Set.t
-    val relocate : S.t Node.Map.t -> t -> t
-
-    val to_condition: t -> (C.t * Node.t option) option
-  end
+let node = Node.create
+
+let identify sigma ~src ~tgt =
+  Sigma.iter2
+    (fun _chunk u v ->
+       match u,v with
+       | Some x , Some y -> F.Subst.add sigma (F.e_var x) (F.e_var y)
+       | _ -> ())
+    src tgt
+
+module E = struct
+  type t = sigma sequence * F.pred
+  let pretty fmt (_seq,p) = Format.fprintf fmt "effect: @[%a@]" F.pp_pred p
+  let get : t -> F.pred = snd
+  let create seq p = seq,p
+
+  let relocate tgt (src,p) =
+    let sigma = Lang.sigma () in
+    identify sigma ~src:src.pre ~tgt:tgt.pre ;
+    identify sigma ~src:src.post ~tgt:tgt.post ;
+    tgt , F.p_subst sigma p
+
+  let reads (seq,_) = Sigma.domain seq.pre
+  let writes ({ pre ; post },_) = Sigma.writes ~pre ~post
+end
 
-  module T :
-  sig
-    type t
-    val pretty : Format.formatter -> t -> unit
-
-    (** Bundle an equation with the sigma sequence that created it. *)
-    val create : S.t Node.Map.t -> F.term -> t
-    val get: t -> F.term
-    val reads : t -> S.domain Node.Map.t
-    val relocate : S.t Node.Map.t -> t -> t
-    val init  : Node.Set.t ->  (S.t Node.Map.t -> F.term) -> t
-    val init' : Node.t -> (S.t -> F.term) -> t
-  end
+module C = struct
+  type t = sigma * F.pred
+  let get = snd
+  let create seq p = seq,p
+  let relocate tgt (src,p) =
+    let sigma = Lang.sigma () in
+    identify sigma ~src ~tgt ;
+    tgt , F.p_subst sigma p
+  let reads (src,_) = Sigma.domain src
+  let equal (s1,p1) (s2,p2) =
+    let sigma = Lang.sigma () in
+    identify sigma ~src:s1 ~tgt:s2 ;
+    F.eqp (F.p_subst sigma p1) p2
+end
 
-  module E : sig
-    type t
-    val pretty: Format.formatter -> t -> unit
-    val create : S.t sequence -> F.pred -> t
-    val get : t -> F.pred
-    val reads : t -> S.domain
-    val writes : t -> S.domain
-    val relocate : S.t sequence -> t -> t
-  end
+module P = struct
+  type t = sigma Node.Map.t * F.pred
+  let pretty fmt (m,f) =
+    Format.fprintf fmt "%a(%a)"
+      F.pp_pred f (Pretty_utils.pp_iter2 Node.Map.iter ~between:",@ " Node.pp (fun _ _ -> ())) m
+  let get = snd
+  let create smap p = smap,p
+
+  let relocate tgt (src,p) =
+    let sigma = Lang.sigma () in
+    Node.Map.iter2
+      (fun n src tgt ->
+         match src,tgt with
+         | Some src , Some tgt -> identify sigma ~src ~tgt
+         | Some _, None ->
+           invalid_arg (Format.asprintf "P.relocate: tgt is smaller than src at %a" Node.pp n)
+         | _ -> ())
+      src tgt ;
+    let tgt = Node.Map.inter (fun _ _ tgt -> tgt) src tgt in
+    tgt , F.p_subst sigma p
+
+  let reads (smap,_) = Node.Map.map (fun _ s -> Sigma.domain s) smap
+  let nodes (smap,_) = Node.Map.fold (fun k _ acc -> Node.Set.add k acc) smap Node.Set.empty
+  let nodes_list (smap,_) = Node.Map.fold (fun k _ acc -> k::acc) smap []
+
+  let to_condition (m,p) =
+    let l = Node.Map.fold (fun k e acc -> (k,e)::acc) m [] in
+    match l with
+    | [] -> Some ((Sigma.create (),p), None)
+    | [n,s] -> Some ((s,p), Some n)
+    | _ -> None
+end
 
-  type cfg
-  val dump_env: name:string -> cfg -> unit
-  val output_dot: out_channel -> ?checks:P.t Bag.t -> cfg -> unit
-
-  val nop : cfg
-  val add_tmpnode: node -> cfg
-  val concat : cfg -> cfg -> cfg
-  val meta : ?stmt:stmt -> ?descr:string -> node -> cfg
-  val goto : node -> node -> cfg
-  val branch : node -> C.t -> node -> node -> cfg
-  val guard : node -> C.t -> node -> cfg
-  val guard' : node -> C.t -> node -> cfg
-  val either : node -> node list -> cfg
-  val implies : node -> (C.t * node) list -> cfg
-  val memory_effect : node -> E.t -> node -> cfg
-  val assume : P.t -> cfg
-  val havoc : node -> memory_effects:node sequence -> node -> cfg
-
-  val compile : ?name:string -> ?mode:mode -> node -> Node.Set.t -> S.domain Node.Map.t ->
-    cfg -> F.pred Node.Map.t * S.t Node.Map.t * Conditions.sequence
+module T = struct
+  type t = sigma Node.Map.t * F.term
 
-end
+  let pretty fmt (m,f) =
+    Format.fprintf fmt "%a(%a)"
+      F.pp_term f (Pretty_utils.pp_iter2 Node.Map.iter ~between:",@ " Node.pp (fun _ _ -> ())) m
 
-module Cfg (S:Sigma) : Cfg with module S = S =
-struct
-
-  module S = S
-
-  module Node : sig
-    type t
-    module Map : Qed.Idxmap.S with type key = t
-    module Set : Qed.Idxset.S with type elt = t
-    module Hashtbl : FCHashtbl.S with type key = t
-    val tag: t -> int
-    val compare: t -> t -> int
-    val equal: t -> t -> bool
-    val pp: Format.formatter -> t -> unit
-    val create: unit -> t
-    val node_internal: unit -> t
-  end
-  = struct
-    type t = int
-    module I = struct type t = int let id x = x end
-    module Map = Qed.Idxmap.Make(I)
-    module Set = Qed.Idxset.Make(I)
-    module Hashtbl = Datatype.Int.Hashtbl
-    let tag = I.id
-    let compare = Datatype.Int.compare
-    let equal = Datatype.Int.equal
-    let pp fmt n =
-      if n>=0 then Format.pp_print_int fmt n
-      else Format.fprintf fmt "int%i" (-n)
-
-    let node_compter = ref (-1)
-
-    let create () =
-      incr node_compter;
-      !node_compter
-
-    let node_internal_compter = ref 0
-
-    let node_internal () =
-      decr node_internal_compter;
-      !node_internal_compter
+  let get = snd
 
-  end
+  let create smap t = smap,t
 
-  let node = Node.create
+  let reads (smap,_) = Node.Map.map (fun _ s -> Sigma.domain s) smap
 
-  let identify sigma ~src ~tgt =
-    S.iter2
-      (fun _chunk u v ->
-         match u,v with
-         | Some x , Some y -> F.Subst.add sigma (F.e_var x) (F.e_var y)
+  let relocate tgt (src,p) =
+    let sigma = Lang.sigma () in
+    Node.Map.iter2
+      (fun _ src tgt ->
+         match src,tgt with
+         | Some src , Some tgt -> identify sigma ~src ~tgt
+         | Some _, None -> invalid_arg "T.relocate: tgt is smaller than src"
          | _ -> ())
-      src tgt
-
-  module E = struct
-    type t = S.t sequence * F.pred
-    let pretty fmt (_seq,p) = Format.fprintf fmt "effect: @[%a@]" F.pp_pred p
-    let get : t -> F.pred = snd
-    let create seq p = seq,p
-
-    let relocate tgt (src,p) =
-      let sigma = Lang.sigma () in
-      identify sigma ~src:src.pre ~tgt:tgt.pre ;
-      identify sigma ~src:src.post ~tgt:tgt.post ;
-      tgt , F.p_subst sigma p
-
-    let reads (seq,_) = S.domain seq.pre
-    let writes (seq,_) = S.writes seq
-  end
+      src tgt ;
+    let tgt = Node.Map.inter (fun _ _ tgt -> tgt) src tgt in
+    tgt , F.e_subst sigma p
+
+  let init node_set f =
+    let node_map = Node.Set.fold (fun x m ->
+        Node.Map.add x (Sigma.create ()) m
+      ) node_set Node.Map.empty
+    in
+    let t = f node_map in
+    (node_map,t)
+
+  let init' node f =
+    let src = Sigma.create () in
+    let t = f src in
+    let node_map =
+      Node.(Map.add node src Map.empty)
+    in
+    (node_map,t)
+end
 
-  module C = struct
-    type t = S.t * F.pred
-    let get = snd
-    let create seq p = seq,p
-    let relocate tgt (src,p) =
-      let sigma = Lang.sigma () in
-      identify sigma ~src ~tgt ;
-      tgt , F.p_subst sigma p
-    let reads (src,_) = S.domain src
-    let equal (s1,p1) (s2,p2) =
-      let sigma = Lang.sigma () in
-      identify sigma ~src:s1 ~tgt:s2 ;
-      F.eqp (F.p_subst sigma p1) p2
+type node = Node.t
+
+type without_bindings = Without_Bindings
+type with_bindings = With_Bindings
+let _ = Without_Bindings
+let _ = With_Bindings
+
+type ('havoc,_) edge =
+  | Goto of node
+  | Branch of C.t * node option * node option
+  | Either of node list
+  | Implies of (C.t * node) list
+  | Effect of E.t * node
+  | Havoc of 'havoc * node
+  | Binding : Passive.t * node -> ('havoc,with_bindings) edge
+  (** Binding used for sigma merging *)
+
+type data =
+  | Meta of stmt option * string option
+
+type ('havoc, 'bindings) env = {
+  succs : ('havoc, 'bindings) edge Node.Map.t;
+  datas : data Bag.t Node.Map.t;
+  (* datas is always included in succs *)
+  assumes : P.t Bag.t;
+  tmpnodes : Node.Set.t; (* node that could be removed *)
+}
+
+
+type pre_env = (node sequence, without_bindings) env
+type restricted_env = (domain, without_bindings) env
+type localised_env = (domain, with_bindings) env
+
+type cfg = pre_env
+
+let iter_succs : type a b. (Node.t -> unit) -> (a,b) edge -> unit = fun f -> function
+  | Goto n2 | Effect(_,n2) | Havoc(_,n2) -> f n2
+  | Branch(_,n2a,n2b) ->
+    let f' = function None -> () | Some x -> f x in
+    f' n2a; f' n2b
+  | Either l -> List.iter f l
+  | Implies l -> List.iter (fun (_,a) -> f a) l
+  | Binding (_,n2) -> f n2
+
+let iter_succs_e f cfg n =
+  match Node.Map.find n cfg.succs with
+  | exception Not_found -> ()
+  | e -> iter_succs f e
+
+let succs : type a b. (a,b) env -> Node.t -> Node.t list =
+  fun cfg n ->
+  match Node.Map.find n cfg.succs with
+  | exception Not_found -> []
+  | Goto n2 | Effect(_,n2) | Havoc(_,n2)
+  | Branch(_,Some n2,None)
+  | Branch(_,None,Some n2) -> [n2]
+  | Binding (_,n2) -> [n2]
+  | Branch(_,Some n1,Some n2) -> [n1;n2]
+  | Branch(_,None,None) -> []
+  | Either l -> l
+  | Implies l -> List.map snd l
+
+let pretty_edge : type a. Format.formatter -> (_,a) edge -> unit = fun fmt edge ->
+  match edge with
+  | Goto(n) -> Format.fprintf fmt "goto(%a)" Node.pp n
+  | Branch(c,n1,n2) -> Format.fprintf fmt "branch(%a,%a,%a)"
+                         Lang.F.pp_pred (C.get c)
+                         (Pretty_utils.pp_opt Node.pp) n1 (Pretty_utils.pp_opt Node.pp) n2
+  | Either l -> Format.fprintf fmt "either(%a)" (Pretty_utils.pp_list ~sep:",@ " Node.pp) l
+  | Implies l -> Format.fprintf fmt "implies(%a)"
+                   (Pretty_utils.pp_list ~sep:",@ " (fun fmt (c,a) ->
+                        Format.fprintf fmt "%a=>%a" Lang.F.pp_pred (C.get c) Node.pp a)) l
+  | Effect(_,n) -> Format.fprintf fmt "effect(%a)" Node.pp n
+  | Havoc(_,n) -> Format.fprintf fmt "havoc(%a)" Node.pp n
+  | Binding(_,n) -> Format.fprintf fmt "binding(%a)" Node.pp n
+
+let pretty_data fmt = function
+  | Meta(s_opt,str_opt) ->
+    Format.fprintf fmt "Meta(%a,%a)"
+      (Pretty_utils.pp_opt ~none:"None" Cil_datatype.Stmt.pretty_sid) s_opt
+      (Pretty_utils.pp_opt ~none:"None" Format.pp_print_string) str_opt
+
+let pretty_env : type a. Format.formatter -> (_,a) env -> unit =
+  fun fmt env ->
+  Context.bind Lang.F.context_pp (Lang.F.env Lang.F.Vars.empty) (fun () ->
+      Format.fprintf fmt
+        "@[<v>@[<3>@[succs:@]@ %a@]@,@[<3>@[datas:@]@ %a@]@,@[<3>@[assumes:@]@ %a@]@]@."
+        (Pretty_utils.pp_iter2 ~between:"->@," ~sep:",@ " Node.Map.iter Node.pp pretty_edge) env.succs
+        (Pretty_utils.pp_iter2 ~between:"->@," ~sep:",@ " Node.Map.iter Node.pp
+           (Pretty_utils.pp_iter Bag.iter pretty_data)) env.datas
+        (Pretty_utils.pp_iter ~sep:",@ " Bag.iter P.pretty) env.assumes
+    ) ()
+
+let dump_edge : type a. node -> Format.formatter -> (_, a) edge -> unit =
+  fun n fmt edge ->
+  let pp_edge ?(label="") n' =
+    Format.fprintf fmt " %a -> %a [ label=\"%s\" ] ;@." Node.pp n Node.pp n' label
+  in
+  begin match edge with
+    | Goto n1 -> pp_edge n1
+    | Branch (_, n1, n2)->
+      Option.iter pp_edge n1;
+      Option.iter pp_edge n2
+    | Either ns -> List.iter pp_edge ns
+    | Implies ns -> List.iter (fun (_,a) -> pp_edge a) ns
+    | Effect (e, n') ->
+      pp_edge ~label:(Format.asprintf "%a" E.pretty e) n'
+    | Havoc (_, n') -> pp_edge ~label:"havoc" n'
+    | Binding (_,n') -> pp_edge ~label:"binding" n'
   end
 
-  module P = struct
-    type t = S.t Node.Map.t * F.pred
-    let pretty fmt (m,f) =
-      Format.fprintf fmt "%a(%a)"
-        F.pp_pred f (Pretty_utils.pp_iter2 Node.Map.iter ~between:",@ " Node.pp (fun _ _ -> ())) m
-    let get = snd
-    let create smap p = smap,p
-
-    let relocate tgt (src,p) =
-      let sigma = Lang.sigma () in
-      Node.Map.iter2
-        (fun n src tgt ->
-           match src,tgt with
-           | Some src , Some tgt -> identify sigma ~src ~tgt
-           | Some _, None ->
-             invalid_arg (Format.asprintf "P.relocate: tgt is smaller than src at %a" Node.pp n)
-           | _ -> ())
-        src tgt ;
-      let tgt = Node.Map.inter (fun _ _ tgt -> tgt) src tgt in
-      tgt , F.p_subst sigma p
-
-    let reads (smap,_) = Node.Map.map (fun _ s -> S.domain s) smap
-    let nodes (smap,_) = Node.Map.fold (fun k _ acc -> Node.Set.add k acc) smap Node.Set.empty
-    let nodes_list (smap,_) = Node.Map.fold (fun k _ acc -> k::acc) smap []
-
-    let to_condition (m,p) =
-      let l = Node.Map.fold (fun k e acc -> (k,e)::acc) m [] in
-      match l with
-      | [] -> Some ((S.create (),p), None)
-      | [n,s] -> Some ((s,p), Some n)
-      | _ -> None
+let dump_node : data Bag.t -> Format.formatter -> node -> unit =
+  fun datas fmt n ->
+  Format.fprintf fmt "  %a [ label=\"%a\n%a\" ] ;@."
+    Node.pp n Node.pp n (Pretty_utils.pp_iter ~sep:"\n" Bag.iter pretty_data) datas
+
+let dump_succ : type a. (_, a) env -> Format.formatter -> node -> (_, a) edge -> unit =
+  fun env fmt n e ->
+  let datas = try Node.Map.find n env.datas with Not_found -> Bag.empty in
+  Format.fprintf fmt "%a\n%a@\n" (dump_node datas) n (dump_edge n) e
+
+let dump_assume : Format.formatter -> P.t -> unit =
+  let count = ref 0 in
+  fun fmt p ->
+    incr count;
+    Format.fprintf fmt "  subgraph cluster_%d {@\n" !count;
+    Format.fprintf fmt "    color=\"palegreen\";@\n";
+    Node.Map.iter
+      (fun n _ -> Format.fprintf fmt "    %a;\n" Node.pp n)
+      (P.reads p);
+    Format.fprintf fmt "    label=\"%a\";" Lang.F.pp_pred (P.get p);
+    Format.fprintf fmt "  }@."
+
+
+let escape fmt = Pretty_utils.ksfprintf (fun s -> String.escaped s) fmt
+
+let output_dot : type a b. out_channel -> ?checks:_ -> (a,b) env -> unit =
+  fun cout ?(checks=Bag.empty) env ->
+  let count = let c = ref max_int in fun () -> decr c; !c in
+  let module E = struct
+    type t = Graph.Graphviz.DotAttributes.edge list
+    let default = []
+    let compare x y = assert (x == y); 0
   end
-
-  module T = struct
-    type t = S.t Node.Map.t * F.term
-
-    let pretty fmt (m,f) =
-      Format.fprintf fmt "%a(%a)"
-        F.pp_term f (Pretty_utils.pp_iter2 Node.Map.iter ~between:",@ " Node.pp (fun _ _ -> ())) m
-
-    let get = snd
-
-    let create smap t = smap,t
-
-    let reads (smap,_) = Node.Map.map (fun _ s -> S.domain s) smap
-
-    let relocate tgt (src,p) =
-      let sigma = Lang.sigma () in
-      Node.Map.iter2
-        (fun _ src tgt ->
-           match src,tgt with
-           | Some src , Some tgt -> identify sigma ~src ~tgt
-           | Some _, None -> invalid_arg "T.relocate: tgt is smaller than src"
-           | _ -> ())
-        src tgt ;
-      let tgt = Node.Map.inter (fun _ _ tgt -> tgt) src tgt in
-      tgt , F.e_subst sigma p
-
-    let init node_set f =
-      let node_map = Node.Set.fold (fun x m ->
-          Node.Map.add x (S.create ()) m
-        ) node_set Node.Map.empty
-      in
-      let t = f node_map in
-      (node_map,t)
-
-    let init' node f =
-      let src = S.create () in
-      let t = f src in
-      let node_map =
-        Node.(Map.add node src Map.empty)
+  in
+  let module V = struct
+    type t =
+      | Node of Node.t
+      | Assume of int * Lang.F.pred
+      | Check of int * Lang.F.pred
+      (* todo better saner comparison *)
+    let tag = function | Node i -> Node.tag i | Assume (i,_) -> i | Check (i,_) -> i
+    let pp fmt = function | Node i -> Node.pp fmt i | Assume (i,_) -> Format.fprintf fmt "ass%i" i
+                          | Check (i,_) -> Format.fprintf fmt "chk%i" i
+    let equal x y = (tag x) = (tag y)
+    let compare x y = Stdlib.compare (tag x) (tag y)
+    let hash x = tag x
+  end in
+  let module G = Graph.Imperative.Digraph.ConcreteBidirectionalLabeled (V)(E) in
+  let module Dot = Graph.Graphviz.Dot(struct
+      let graph_attributes _g = [`Fontname "fixed"]
+      let default_vertex_attributes _g = (* [`Shape `Point] *) [`Shape `Circle]
+      let vertex_name v = Format.asprintf "cp%a" V.pp  v
+      let vertex_attributes  = function
+        | V.Node n -> [`Label (escape "%a" Node.pp n)]
+        | V.Assume (_,p) -> [`Style `Dashed; `Label (escape "%a" Lang.F.pp_pred p)]
+        | V.Check (_,p) -> [`Style `Dotted; `Label (escape "%a" Lang.F.pp_pred p)]
+      let get_subgraph _ = None
+      let default_edge_attributes _g = []
+      let edge_attributes ((_,e,_):G.E.t) : Graph.Graphviz.DotAttributes.edge list = e
+      include G
+    end) in
+  let g = G.create () in
+  let add_edge n1 l n2 =  G.add_edge_e g (V.Node n1,l,V.Node n2) in
+  let add_edges : type a b. Node.t -> (a,b) edge -> unit = fun n1 -> function
+    | Goto n2 -> add_edge n1 [] n2
+    | Branch((_,c),n2,n2') ->
+      let aux s = function
+        | None -> ()
+        | Some n -> add_edge n1 [`Label (escape "%s%a" s Lang.F.pp_pred c)] n
       in
-      (node_map,t)
+      aux "" n2; aux "!" n2'
+    | Either l -> List.iter (add_edge n1 []) l
+    | Implies l ->
+      List.iter (fun (c,n) -> add_edge n1 [`Label (escape "%a" Lang.F.pp_pred (C.get c))] n) l
+    | Effect ((_,e),n2) ->
+      add_edge n1 [`Label (escape "%a" Lang.F.pp_pred e)] n2
+    | Havoc (_,n2) -> add_edge n1 [`Label (escape "havoc")] n2
+    | Binding (_,n2) -> add_edge n1 [`Label (escape "binding")] n2
+  in
+  Node.Map.iter add_edges env.succs;
+  (* assumes *)
+  Bag.iter (fun (m,p) ->
+      let n1 = V.Assume(count (), p) in
+      let assume_label = [`Style `Dashed ] in
+      Node.Map.iter (fun n2 _ -> G.add_edge_e g (n1,assume_label,V.Node n2)) m
+    ) env.assumes;
+  (* checks *)
+  Bag.iter (fun (m,p) ->
+      let n1 = V.Check(count (), p) in
+      let label = [`Style `Dotted ] in
+      Node.Map.iter (fun n2 _ -> G.add_edge_e g (V.Node n2,label,n1)) m
+    ) checks;
+  Dot.output_graph cout g
+
+let dump_env : type a. name:string -> (_, a) env -> unit = fun ~name env ->
+  let file = (Filename.get_temp_dir_name ()) ^ "/cfg_" ^ name in
+  let fout = open_out (file ^ ".dot") in
+  if false then begin
+    let out = Format.formatter_of_out_channel fout in
+    Format.fprintf out "digraph %s {@\n" name;
+    Format.fprintf out "  rankdir = TB ;@\n";
+    Format.fprintf out "  node [ style = filled, shape = circle ] ;@\n";
+    Node.Map.iter (dump_succ env out) env.succs;
+    Bag.iter (dump_assume out) env.assumes;
+    Format.fprintf out "}@.";
   end
-
-  type node = Node.t
-
-  type without_bindings = Without_Bindings
-  type with_bindings = With_Bindings
-  let _ = Without_Bindings
-  let _ = With_Bindings
-
-  type ('havoc,_) edge =
-    | Goto of node
-    | Branch of C.t * node option * node option
-    | Either of node list
-    | Implies of (C.t * node) list
-    | Effect of E.t * node
-    | Havoc of 'havoc * node
-    | Binding : Passive.t * node -> ('havoc,with_bindings) edge
-    (** Binding used for sigma merging *)
-
-  type data =
-    | Meta of stmt option * string option
-
-  type ('havoc, 'bindings) env = {
-    succs : ('havoc, 'bindings) edge Node.Map.t;
-    datas : data Bag.t Node.Map.t;
-    (* datas is always included in succs *)
-    assumes : P.t Bag.t;
-    tmpnodes : Node.Set.t; (* node that could be removed *)
+  else begin
+    output_dot fout env;
+  end;
+  close_out fout;
+  ignore (Sys.command
+            (Printf.sprintf "dot -Tpdf %s.dot > %s.pdf" file file));
+  Wp_parameters.debug ~dkey:dumpkey "Saving dump %s into %s.pdf" name file
+
+let env_union env1 env2 =
+  {
+    succs = Node.Map.union
+        (fun _ _v1 _v2 -> invalid_arg "A node has more than one successor")
+        env1.succs env2.succs;
+    datas = Node.Map.union (fun _ -> Bag.concat) env1.datas env2.datas;
+    assumes = Bag.concat env1.assumes env2.assumes;
+    tmpnodes = Node.Set.union env1.tmpnodes env2.tmpnodes;
   }
 
+let new_env ?(succs=Node.Map.empty) ?(datas=Node.Map.empty) ?(assumes=Bag.empty)
+    ?(tmpnodes=Node.Set.empty) () =
+  {succs; datas; assumes; tmpnodes}
+
+let nop = new_env ()
+
+let add_tmpnode node = new_env ~tmpnodes:(Node.Set.singleton node) ()
+
+let concat a b = env_union a b
+
+let meta ?stmt ?descr n =
+  let data = Meta(stmt,descr) in
+  new_env ~datas:(Node.Map.add n (Bag.elt data) (Node.Map.empty)) ()
+
+let edge n e =
+  new_env ~succs:(Node.Map.add n e (Node.Map.empty)) ()
+
+let goto node_orig node_target =
+  edge node_orig (Goto(node_target))
+
+let branch node_orig predicate node_target_then node_target_else =
+  edge node_orig (Branch(predicate,
+                         Some node_target_then,
+                         Some node_target_else))
+
+let guard node_orig predicate node_target_then =
+  edge node_orig (Branch(predicate,
+                         Some node_target_then,
+                         None))
+
+let guard' node_orig predicate node_target_else =
+  edge node_orig (Branch(predicate,
+                         None,
+                         Some node_target_else
+                        ))
+
+let either node = function
+  | [] -> nop
+  | [dest] -> goto node dest
+  | node_list -> edge node (Either(node_list))
+
+let implies node = function
+  | [] -> nop
+  | [g,dest] -> guard node g dest
+  | node_list -> edge node (Implies(node_list))
+
+let memory_effect node1 e node2 =
+  edge node1 (Effect(e, node2))
+
+let assume (predicate:P.t) =
+  if F.is_ptrue (P.get predicate) = Qed.Logic.Yes
+  then nop
+  else new_env ~assumes:(Bag.elt predicate) ()
+
+let havoc node1 ~memory_effects:node_seq node2 =
+  edge node1 (Havoc(node_seq,node2))
+
+let option_bind ~f = function
+  | None -> None
+  | Some x -> f x
+
+let union_opt_or union d1 d2 =
+  match d1, d2 with
+  | Some d1, Some d2 -> Some (union d1 d2)
+  | (Some  _ as d), None | None, (Some _ as d) -> d
+  | None, None -> None
+
+let union_opt_and union d1 d2 =
+  match d1, d2 with
+  | Some d1, Some d2 -> Some (union d1 d2)
+  | _ -> None
+
+let add_only_if_alive union d1 = function
+  | None -> None
+  | Some d2 -> Some (union d1 d2)
+
+(** return None when post is not accessible from this node *)
+let rec effects : type a.  (_,a) env -> node -> node -> domain option =
+  fun env post node ->
+  if node = post
+  then Some Sigma.empty
+  else
+    match Node.Map.find node env.succs with
+    | exception Not_found -> None
+    | Goto (node2) ->
+      effects env post node2
+    | Branch (_, node2, node3) ->
+      union_opt_or Sigma.union
+        (option_bind ~f:(effects env post) node2)
+        (option_bind ~f:(effects env post) node3)
+    | Either (l) ->
+      (List.fold_left
+         (fun acc node2 -> union_opt_or Sigma.union
+             acc (effects env post node2))
+         None l)
+    | Implies (l) ->
+      (List.fold_left
+         (fun acc (_,node2) -> union_opt_or Sigma.union
+             acc (effects env post node2))
+         None l)
+    | Effect (memory_effect , node2) ->
+      add_only_if_alive Sigma.union
+        (E.writes memory_effect)
+        (effects env post node2)
+    | Havoc (m, node2) ->
+      union_opt_and Sigma.union
+        (effects env m.post m.pre)
+        (effects env post node2)
+    | Binding (_,node2) ->
+      effects env post node2
+
+(** restrict a cfg to the nodes accessible from the pre post given,
+    and compute havoc effect *)
+let restrict (cfg:pre_env) pre posts : restricted_env =
+  let rec walk acc node : restricted_env option =
+    if Node.Map.mem node acc.succs then Some acc
+    else
+      let new_env edge = new_env ~succs:(Node.Map.add node edge (Node.Map.empty)) () in
+      let r = match Node.Map.find node cfg.succs with
+        | exception Not_found -> None
+        | (Goto (node2) | Effect (_ , node2)) as edge ->
+          union_opt_and env_union
+            (Some (new_env edge))
+            (walk acc node2)
+        | Branch (pred, node2, node3) ->
+          (* it is important to visit all the childrens *)
+          let f acc node =
+            match option_bind ~f:(walk acc) node with
+            | None -> None, acc
+            | Some acc -> node, acc in
+          let node2, acc = f acc node2 in
+          let node3, acc = f acc node3 in
+          if node2 = None && node3 = None then None
+          else Some (env_union acc (new_env (Branch(pred, node2, node3))))
+        | Either (l) ->
+          let acc,l = List.fold_left
+              (fun ((acc,l) as old) node2 ->
+                 match walk acc node2 with
+                 | None -> old
+                 | Some acc -> (acc,node2::l))
+              (acc,[]) l in
+          if l = [] then None
+          else Some (env_union acc (new_env (Either (List.rev l))))
+        | Implies (l) ->
+          let acc,l = List.fold_left
+              (fun ((acc,l) as old) ((_,node2) as e) ->
+                 match walk acc node2 with
+                 | None -> old
+                 | Some acc -> (acc,e::l))
+              (acc,[]) l in
+          if l = [] then None
+          else Some (env_union acc (new_env (Implies (List.rev l))))
+        | Havoc (m, node2) ->
+          match effects cfg m.post m.pre with
+          | None -> None
+          | Some eff ->
+            union_opt_and env_union
+              (Some (new_env (Havoc(eff,node2))))
+              (walk acc node2)
+      in
+      if Node.Set.mem node posts && r = None
+      then Some acc
+      else r
+  in
+  match walk (new_env ()) pre with
+  | None -> (new_env ())
+  | Some acc ->
+    { succs = acc.succs;
+      datas = Node.Map.inter (fun _ _ v -> v) acc.succs cfg.datas;
+      assumes = Bag.filter (fun (seq,_) ->
+          Node.Map.subset (fun _ _ _ -> true)
+            seq acc.succs) cfg.assumes;
+      tmpnodes = cfg.tmpnodes;
+    }
 
-  type pre_env = (node sequence, without_bindings) env
-  type restricted_env = (S.domain, without_bindings) env
-  type localised_env = (S.domain, with_bindings) env
-
-  type cfg = pre_env
-
-  let iter_succs : type a b. (Node.t -> unit) -> (a,b) edge -> unit = fun f -> function
-    | Goto n2 | Effect(_,n2) | Havoc(_,n2) -> f n2
-    | Branch(_,n2a,n2b) ->
-      let f' = function None -> () | Some x -> f x in
-      f' n2a; f' n2b
-    | Either l -> List.iter f l
-    | Implies l -> List.iter (fun (_,a) -> f a) l
-    | Binding (_,n2) -> f n2
-
-  let iter_succs_e f cfg n =
-    match Node.Map.find n cfg.succs with
-    | exception Not_found -> ()
-    | e -> iter_succs f e
-
-  let succs : type a b. (a,b) env -> Node.t -> Node.t list =
-    fun cfg n ->
-    match Node.Map.find n cfg.succs with
-    | exception Not_found -> []
-    | Goto n2 | Effect(_,n2) | Havoc(_,n2)
-    | Branch(_,Some n2,None)
-    | Branch(_,None,Some n2) -> [n2]
-    | Binding (_,n2) -> [n2]
-    | Branch(_,Some n1,Some n2) -> [n1;n2]
-    | Branch(_,None,None) -> []
-    | Either l -> l
-    | Implies l -> List.map snd l
-
-  let pretty_edge : type a. Format.formatter -> (_,a) edge -> unit = fun fmt edge ->
-    match edge with
-    | Goto(n) -> Format.fprintf fmt "goto(%a)" Node.pp n
-    | Branch(c,n1,n2) -> Format.fprintf fmt "branch(%a,%a,%a)"
-                           Lang.F.pp_pred (C.get c)
-                           (Pretty_utils.pp_opt Node.pp) n1 (Pretty_utils.pp_opt Node.pp) n2
-    | Either l -> Format.fprintf fmt "either(%a)" (Pretty_utils.pp_list ~sep:",@ " Node.pp) l
-    | Implies l -> Format.fprintf fmt "implies(%a)"
-                     (Pretty_utils.pp_list ~sep:",@ " (fun fmt (c,a) ->
-                          Format.fprintf fmt "%a=>%a" Lang.F.pp_pred (C.get c) Node.pp a)) l
-    | Effect(_,n) -> Format.fprintf fmt "effect(%a)" Node.pp n
-    | Havoc(_,n) -> Format.fprintf fmt "havoc(%a)" Node.pp n
-    | Binding(_,n) -> Format.fprintf fmt "binding(%a)" Node.pp n
-
-  let pretty_data fmt = function
-    | Meta(s_opt,str_opt) ->
-      Format.fprintf fmt "Meta(%a,%a)"
-        (Pretty_utils.pp_opt ~none:"None" Cil_datatype.Stmt.pretty_sid) s_opt
-        (Pretty_utils.pp_opt ~none:"None" Format.pp_print_string) str_opt
-
-  let pretty_env : type a. Format.formatter -> (_,a) env -> unit =
-    fun fmt env ->
-    Context.bind Lang.F.context_pp (Lang.F.env Lang.F.Vars.empty) (fun () ->
-        Format.fprintf fmt
-          "@[<v>@[<3>@[succs:@]@ %a@]@,@[<3>@[datas:@]@ %a@]@,@[<3>@[assumes:@]@ %a@]@]@."
-          (Pretty_utils.pp_iter2 ~between:"->@," ~sep:",@ " Node.Map.iter Node.pp pretty_edge) env.succs
-          (Pretty_utils.pp_iter2 ~between:"->@," ~sep:",@ " Node.Map.iter Node.pp
-             (Pretty_utils.pp_iter Bag.iter pretty_data)) env.datas
-          (Pretty_utils.pp_iter ~sep:",@ " Bag.iter P.pretty) env.assumes
-      ) ()
-
-  let dump_edge : type a. node -> Format.formatter -> (_, a) edge -> unit =
-    fun n fmt edge ->
-    let pp_edge ?(label="") n' =
-      Format.fprintf fmt " %a -> %a [ label=\"%s\" ] ;@." Node.pp n Node.pp n' label
+(** succ is decreasing for this order *)
+let topological (type a) (type b) (cfg:(a,b) env) =
+  let module G = struct
+    type t = (a,b) env
+    module V = struct let hash = Hashtbl.hash include Node end
+    let iter_vertex f cfg =
+      let h = Node.Hashtbl.create 10 in
+      let replace n = Node.Hashtbl.replace h n () in
+      Node.Map.iter (fun k _ -> replace k; iter_succs_e replace cfg k) cfg.succs;
+      Node.Hashtbl.iter (fun k () -> f k) h
+    let iter_succ = iter_succs_e
+  end in
+  let module T = Graph.Topological.Make(G) in
+  let h  = Node.Hashtbl.create 10 in
+  let h' = Datatype.Int.Hashtbl.create 10 in
+  let c = ref (-1) in
+  let l = ref [] in
+  T.iter (fun n -> l := n::!l; incr c; Node.Hashtbl.add h n !c; Datatype.Int.Hashtbl.add h' !c n) cfg;
+  h,h',List.rev !l
+
+(** topo_list: elements in topological order
+    topo_order: post-order mapping
+    nb: number of elements *)
+let idoms topo_list topo_order nb ~pred ~is_after =
+  let a = Array.make nb (-1) in
+  let iter n =
+    let first,preds = match pred n with
+      | [] -> topo_order n, []
+      | f::p -> topo_order f, List.map topo_order p
     in
-    begin match edge with
-      | Goto n1 -> pp_edge n1
-      | Branch (_, n1, n2)->
-        Option.iter pp_edge n1;
-        Option.iter pp_edge n2
-      | Either ns -> List.iter pp_edge ns
-      | Implies ns -> List.iter (fun (_,a) -> pp_edge a) ns
-      | Effect (e, n') ->
-        pp_edge ~label:(Format.asprintf "%a" E.pretty e) n'
-      | Havoc (_, n') -> pp_edge ~label:"havoc" n'
-      | Binding (_,n') -> pp_edge ~label:"binding" n'
-    end
-
-  let dump_node : data Bag.t -> Format.formatter -> node -> unit =
-    fun datas fmt n ->
-    Format.fprintf fmt "  %a [ label=\"%a\n%a\" ] ;@."
-      Node.pp n Node.pp n (Pretty_utils.pp_iter ~sep:"\n" Bag.iter pretty_data) datas
-
-  let dump_succ : type a. (_, a) env -> Format.formatter -> node -> (_, a) edge -> unit =
-    fun env fmt n e ->
-    let datas = try Node.Map.find n env.datas with Not_found -> Bag.empty in
-    Format.fprintf fmt "%a\n%a@\n" (dump_node datas) n (dump_edge n) e
-
-  let dump_assume : Format.formatter -> P.t -> unit =
-    let count = ref 0 in
-    fun fmt p ->
-      incr count;
-      Format.fprintf fmt "  subgraph cluster_%d {@\n" !count;
-      Format.fprintf fmt "    color=\"palegreen\";@\n";
-      Node.Map.iter
-        (fun n _ -> Format.fprintf fmt "    %a;\n" Node.pp n)
-        (P.reads p);
-      Format.fprintf fmt "    label=\"%a\";" Lang.F.pp_pred (P.get p);
-      Format.fprintf fmt "  }@."
-
-
-  let escape fmt = Pretty_utils.ksfprintf (fun s -> String.escaped s) fmt
-
-  let output_dot : type a b. out_channel -> ?checks:_ -> (a,b) env -> unit =
-    fun cout ?(checks=Bag.empty) env ->
-    let count = let c = ref max_int in fun () -> decr c; !c in
-    let module E = struct
-      type t = Graph.Graphviz.DotAttributes.edge list
-      let default = []
-      let compare x y = assert (x == y); 0
-    end
+    let rec find_common n1 n2 =
+      if n1 = n2 then n1
+      else if is_after n1 n2 then find_common a.(n1) n2
+      else find_common n1 a.(n2) in
+    let idom = List.fold_left find_common first preds in
+    a.(topo_order n) <- idom
+  in
+  List.iter iter topo_list;
+  a
+
+let find_def ~def x t = try Node.Map.find x t with Not_found -> def
+
+let rec remove_dumb_gotos (env:restricted_env) : Node.t Node.Map.t * restricted_env =
+  let add_map m acc = Node.Map.fold (fun n _ acc -> Node.Set.add n acc) m acc in
+  let used_nodes =
+    Bag.fold_left (fun acc p -> add_map (P.reads p) acc) Node.Set.empty env.assumes
+  in
+  let used_nodes = add_map env.datas used_nodes in
+  let how_many_preds = Node.Hashtbl.create 10 in
+  let incr_how_many_preds n =
+    Node.Hashtbl.replace how_many_preds n (succ (Node.Hashtbl.find_def how_many_preds n 0))
+  in
+  let subst =
+    Node.Map.fold (fun n e acc ->
+        iter_succs incr_how_many_preds e;
+        match (e:(_,without_bindings) edge) with
+        | Goto n' when not (Node.Set.mem n used_nodes) ->
+          Node.Map.add n n' acc
+        | Goto _
+        | Branch (_,_,_)
+        | Either _
+        | Implies _
+        | Effect (_,_)
+        | Havoc (_,_) -> acc)
+      env.succs Node.Map.empty
+  in
+  let subst =
+    let rec compress n =
+      match Node.Map.find n subst with
+      | exception Not_found -> n
+      | n -> compress n
     in
-    let module V = struct
-      type t =
-        | Node of Node.t
-        | Assume of int * Lang.F.pred
-        | Check of int * Lang.F.pred
-        (* todo better saner comparison *)
-      let tag = function | Node i -> Node.tag i | Assume (i,_) -> i | Check (i,_) -> i
-      let pp fmt = function | Node i -> Node.pp fmt i | Assume (i,_) -> Format.fprintf fmt "ass%i" i
-                            | Check (i,_) -> Format.fprintf fmt "chk%i" i
-      let equal x y = (tag x) = (tag y)
-      let compare x y = Stdlib.compare (tag x) (tag y)
-      let hash x = tag x
-    end in
-    let module G = Graph.Imperative.Digraph.ConcreteBidirectionalLabeled (V)(E) in
-    let module Dot = Graph.Graphviz.Dot(struct
-        let graph_attributes _g = [`Fontname "fixed"]
-        let default_vertex_attributes _g = (* [`Shape `Point] *) [`Shape `Circle]
-        let vertex_name v = Format.asprintf "cp%a" V.pp  v
-        let vertex_attributes  = function
-          | V.Node n -> [`Label (escape "%a" Node.pp n)]
-          | V.Assume (_,p) -> [`Style `Dashed; `Label (escape "%a" Lang.F.pp_pred p)]
-          | V.Check (_,p) -> [`Style `Dotted; `Label (escape "%a" Lang.F.pp_pred p)]
-        let get_subgraph _ = None
-        let default_edge_attributes _g = []
-        let edge_attributes ((_,e,_):G.E.t) : Graph.Graphviz.DotAttributes.edge list = e
-        include G
-      end) in
-    let g = G.create () in
-    let add_edge n1 l n2 =  G.add_edge_e g (V.Node n1,l,V.Node n2) in
-    let add_edges : type a b. Node.t -> (a,b) edge -> unit = fun n1 -> function
-      | Goto n2 -> add_edge n1 [] n2
-      | Branch((_,c),n2,n2') ->
-        let aux s = function
-          | None -> ()
-          | Some n -> add_edge n1 [`Label (escape "%s%a" s Lang.F.pp_pred c)] n
-        in
-        aux "" n2; aux "!" n2'
-      | Either l -> List.iter (add_edge n1 []) l
+    Node.Map.map (fun _ n' -> compress n') subst
+  in
+  let find n = find_def ~def:n n subst in
+  (* detect either that could be transformed in branch *)
+  let to_remove = Node.Hashtbl.create 10 in
+  Node.Map.iter (fun _ e ->
+      match (e:(_,without_bindings) edge) with
+      | Either [a;b] when Node.Hashtbl.find how_many_preds a = 1 &&
+                          Node.Hashtbl.find how_many_preds b = 1 &&
+                          not (Node.Set.mem a used_nodes) &&
+                          not (Node.Set.mem b used_nodes) &&
+                          Node.Set.mem a env.tmpnodes &&
+                          Node.Set.mem b env.tmpnodes &&
+                          true
+        ->
+        begin
+          let find_opt k m =
+            match Node.Map.find k m with
+            | exception Not_found -> None
+            | v -> Some v
+          in
+          match find_opt a env.succs, find_opt b env.succs with
+          | Some Branch(c,Some n1, None), Some Branch(c',None, Some n2)
+          | Some Branch(c,None, Some n2), Some Branch(c',Some n1,None) when C.equal c c' ->
+            let n1 = find n1 in
+            let n2 = find n2 in
+            let br = Branch(c,Some n1, Some n2) in
+            Node.Hashtbl.add to_remove a br;
+            Node.Hashtbl.add to_remove b br
+          | _ -> ()
+        end
+      | Goto _
+      | Branch (_,_,_)
+      | Effect (_,_)
+      | Either _
+      | Implies _
+      | Havoc (_,_) -> ()
+    ) env.succs;
+  (* substitute and remove *)
+  let succs = Node.Map.mapq (fun n e ->
+      match (e:(_,without_bindings) edge) with
+      | _ when Node.Hashtbl.mem to_remove n -> None
+      | Goto _ when not (Node.Set.mem n used_nodes) -> None
+      | Goto n' ->
+        let n'' = find n' in
+        if Node.equal n' n'' then Some e
+        else Some (Goto n'')
+      | Branch (c,n1,n2) ->
+        let n1' = Option.map find n1 in
+        let n2' = Option.map find n2 in
+        if Option.equal Node.equal n1 n1' && Option.equal Node.equal n2 n2'
+        then Some e
+        else Some (Branch(c,n1',n2'))
+      | Either l ->
+        let l' = List.map find l in
+        let l' = List.sort_uniq Node.compare l' in
+        begin match l' with
+          | [] -> assert false (* absurd: Either after restricted has at least one successor *)
+          | [a] -> Some (Goto a)
+          | [a;_] when Node.Hashtbl.mem to_remove a ->
+            let br = Node.Hashtbl.find to_remove a in
+            Some br
+          | l' -> Some (Either l')
+        end
       | Implies l ->
-        List.iter (fun (c,n) -> add_edge n1 [`Label (escape "%a" Lang.F.pp_pred (C.get c))] n) l
-      | Effect ((_,e),n2) ->
-        add_edge n1 [`Label (escape "%a" Lang.F.pp_pred e)] n2
-      | Havoc (_,n2) -> add_edge n1 [`Label (escape "havoc")] n2
-      | Binding (_,n2) -> add_edge n1 [`Label (escape "binding")] n2
+        let l' = List.map (fun (g,n) -> (g,find n)) l in
+        Some (Implies l')
+      | Effect (ef,n') ->
+        let n'' = find n' in
+        if Node.equal n' n'' then Some e
+        else Some (Effect(ef,n''))
+      | Havoc (h,n') ->
+        let n'' = find n' in
+        if Node.equal n' n'' then Some e
+        else Some (Havoc(h,n''))
+    )
+      env.succs
+  in
+  let env = {env with succs} in
+  if Node.Map.is_empty subst
+  then subst, env
+  else
+    let subst', env = remove_dumb_gotos env in
+    let subst = Node.Map.map (fun _ n' -> find_def ~def:n' n' subst') subst in
+    Node.Map.merge (fun _ a b ->
+        match a, b with
+        | Some _, Some _ -> assert false (* the elements are remove in the new env *)
+        | Some x, None | None, Some x -> Some x
+        | None, None -> assert false
+      ) subst subst', env
+
+let allocate domain sigma =
+  Domain.iter (fun chunk -> ignore (Sigma.get sigma chunk)) domain
+
+let domains (env : restricted_env) reads pre : localised_env * sigma Node.Map.t =
+  let visited = ref Node.Map.empty in
+  let new_succs = ref Node.Map.empty in
+  let add_edge node edge = new_succs := Node.Map.add node edge !new_succs in
+  let add_binding_edge n (p: Passive.t) =
+    if Passive.is_empty p then n
+    else
+      let n' = Node.node_internal () in
+      add_edge n' (Binding(p,n));
+      n'
+  in
+  let rec aux node : sigma =
+    try Node.Map.find node !visited
+    with Not_found ->
+      let dom = find_def ~def:Sigma.empty node reads in
+      let ret =
+        match Node.Map.find node env.succs with
+        | exception Not_found ->
+          (* posts node *)
+          let s1 = Sigma.create () in
+          allocate dom s1;
+          s1
+        | Goto (node2) ->
+          let s1 = Sigma.copy (aux node2) in
+          allocate dom s1;
+          add_edge node (Goto node2);
+          s1
+        | Branch (pred, node2, node3) ->
+          let dom = (Sigma.union (C.reads pred) dom) in
+          begin match node2, node3 with
+            | (None, Some next) | (Some next, None) ->
+              let s1 = Sigma.copy (aux next) in
+              allocate dom s1;
+              let pred = C.relocate s1 pred in
+              add_edge node (Branch(pred,node2,node3));
+              s1
+            | Some node2, Some node3 ->
+              let s2 = aux node2 in
+              let s3 = aux node3 in
+              let s1,p2,p3 = Sigma.merge s2 s3 in
+              allocate dom s1;
+              let node2' = add_binding_edge node2 p2 in
+              let node3' = add_binding_edge node3 p3 in
+              let pred = C.relocate s1 pred in
+              add_edge node (Branch(pred,Some node2',Some node3'));
+              s1
+            | _ -> assert false
+          end
+        | Either (l) ->
+          let s1, pl = Sigma.merge_list (List.map aux l) in
+          allocate dom s1;
+          let l = List.map2 add_binding_edge l pl in
+          add_edge node (Either l);
+          s1
+        | Implies (l) ->
+          let dom =
+            List.fold_left (fun acc (c,_) -> Sigma.union (C.reads c) acc) dom l
+          in
+          let s1, pl = Sigma.merge_list (List.map (fun (_,n) -> aux n) l) in
+          allocate dom s1;
+          let l = List.map2 (fun (c,a) b ->
+              let a = add_binding_edge a b in
+              let c = C.relocate s1 c in
+              (c,a)) l pl
+          in
+          add_edge node (Implies l);
+          s1
+        | Effect (memory_effect , node2) ->
+          let s2 = aux node2 in
+          let s1 = Sigma.remove_chunks s2 (E.writes memory_effect) in
+          allocate dom s1;
+          allocate (E.reads memory_effect) s1;
+          let memory_effect = E.relocate {pre=s1;post=s2} memory_effect in
+          add_edge node (Effect(memory_effect,node2));
+          s1
+        | Havoc (eff, node2) ->
+          let s2 = aux node2 in
+          let s1 = Sigma.havoc s2 eff in
+          allocate dom s1;
+          add_edge node (Havoc(eff,node2));
+          s1
+      in
+      visited := Node.Map.add node ret !visited;
+      ret
+  in
+  ignore (aux pre);
+  let sigmas = !visited in
+  let new_env =
+    {succs = !new_succs;
+     datas = env.datas;
+     assumes = Bag.map (fun e -> P.relocate sigmas e) env.assumes;
+     tmpnodes = env.tmpnodes;
+    } in
+  new_env, sigmas
+
+let compute_preds env =
+  let h = Node.Hashtbl.create 10 in
+  let add = Node.Hashtbl.add h in
+  Node.Map.iter (fun n s ->
+      match s with
+      | Goto n1 | Havoc (_, n1) | Effect (_,n1) | Binding (_,n1) -> add n1 n
+      | Branch (_,Some n1,Some n2) ->
+        add n1 n;
+        add n2 n
+      | Branch(_,Some n1,None) -> add n1 n
+      | Branch(_,None,Some n1) -> add n1 n
+      | Branch(_,None,None) -> ()
+      | Either l -> List.iter (fun n1 -> add n1 n) l
+      | Implies l -> List.iter (fun (_,n1) -> add n1 n) l
+    ) env.succs;
+  h
+
+let to_sequence_bool ~mode pre posts env : Conditions.sequence * F.pred Node.Map.t =
+  let preds = Node.Hashtbl.create 10 in
+  let access n = Node.Hashtbl.memo preds n
+      (fun _ ->
+         let v = F.fresh ~basename:"node"
+             (get_pool ()) Qed.Logic.Bool in
+         F.p_bool (F.e_var v))
+  in
+  let (!.) c = (Conditions.sequence [Conditions.step c]) in
+  let have_access n = !. (Conditions.Have (access n)) in
+  let add_cond ?descr ?stmt f cond =
+    Conditions.append (Conditions.sequence [Conditions.step ?descr ?stmt cond]) f
+  in
+  let either = function
+    | [] -> !. (Conditions.Have F.p_false)
+    | [a] -> a
+    | l -> !. (Conditions.Either l)
+  in
+  let f = Conditions.empty in
+  (* The start state is accessible *)
+  let pre = Conditions.Have (access pre) in
+  let f = add_cond f pre in
+  (* The posts state are accessible *)
+  let f = Node.Set.fold
+      (fun n f -> add_cond f (Conditions.Have (access n)))
+      posts f in
+  (* The assumes are true if all their nodes are accessible *)
+  let f =
+    Bag.fold_left (fun f p ->
+        let nodes_are_accessible =
+          Node.Map.fold (fun n _ acc -> F.p_and (access n) acc)
+            (P.reads p) F.p_true in
+        let f' = F.p_imply nodes_are_accessible (P.get p) in
+        add_cond f (Conditions.Have f')
+      ) f env.assumes in
+
+  (* compute predecessors *)
+  let to_sequence_basic_backward f =
+    let predecessors = Node.Map.fold (fun n s acc ->
+        let add acc n' p =
+          Node.Map.change (fun _ (n,p) -> function
+              | None -> Some (Node.Map.add n p Node.Map.empty)
+              | Some s -> Some (Node.Map.add n p s)) n' (n,p) acc
+        in
+        match s with
+        | Goto n' | Havoc (_, n') -> add acc n' F.p_true
+        | Branch (c,Some n1,Some n2) ->
+          let c = P.get c in
+          add (add acc n1 c) n2 (F.p_not c)
+        | Branch(c,Some n1,None) -> add acc n1 (P.get c)
+        | Branch(c,None,Some n1) -> add acc n1 (F.p_not (P.get c))
+        | Branch(_,None,None) -> acc
+        | Either l -> List.fold_left (fun acc e -> add acc e F.p_true) acc l
+        | Implies l -> List.fold_left (fun acc (c,e) -> add acc e (P.get c)) acc l
+        | Effect (e,n') -> add acc n' (E.get e)
+        | Binding (b,n') ->
+          let b = F.p_conj (Passive.conditions b (fun _ -> true)) in
+          add acc n' b
+      ) env.succs Node.Map.empty
     in
-    Node.Map.iter add_edges env.succs;
-    (* assumes *)
-    Bag.iter (fun (m,p) ->
-        let n1 = V.Assume(count (), p) in
-        let assume_label = [`Style `Dashed ] in
-        Node.Map.iter (fun n2 _ -> G.add_edge_e g (n1,assume_label,V.Node n2)) m
-      ) env.assumes;
-    (* checks *)
-    Bag.iter (fun (m,p) ->
-        let n1 = V.Check(count (), p) in
-        let label = [`Style `Dotted ] in
-        Node.Map.iter (fun n2 _ -> G.add_edge_e g (V.Node n2,label,n1)) m
-      ) checks;
-    Dot.output_graph cout g
-
-  let dump_env : type a. name:string -> (_, a) env -> unit = fun ~name env ->
-    let file = (Filename.get_temp_dir_name ()) ^ "/cfg_" ^ name in
-    let fout = open_out (file ^ ".dot") in
-    if false then begin
-      let out = Format.formatter_of_out_channel fout in
-      Format.fprintf out "digraph %s {@\n" name;
-      Format.fprintf out "  rankdir = TB ;@\n";
-      Format.fprintf out "  node [ style = filled, shape = circle ] ;@\n";
-      Node.Map.iter (dump_succ env out) env.succs;
-      Bag.iter (dump_assume out) env.assumes;
-      Format.fprintf out "}@.";
-    end
-    else begin
-      output_dot fout env;
-    end;
-    close_out fout;
-    ignore (Sys.command
-              (Printf.sprintf "dot -Tpdf %s.dot > %s.pdf" file file));
-    Wp_parameters.debug ~dkey:dumpkey "Saving dump %s into %s.pdf" name file
-
-  let env_union env1 env2 =
-    {
-      succs = Node.Map.union
-          (fun _ _v1 _v2 -> invalid_arg "A node has more than one successor")
-          env1.succs env2.succs;
-      datas = Node.Map.union (fun _ -> Bag.concat) env1.datas env2.datas;
-      assumes = Bag.concat env1.assumes env2.assumes;
-      tmpnodes = Node.Set.union env1.tmpnodes env2.tmpnodes;
-    }
-
-  let new_env ?(succs=Node.Map.empty) ?(datas=Node.Map.empty) ?(assumes=Bag.empty)
-      ?(tmpnodes=Node.Set.empty) () =
-    {succs; datas; assumes; tmpnodes}
-
-  let nop = new_env ()
-
-  let add_tmpnode node = new_env ~tmpnodes:(Node.Set.singleton node) ()
-
-  let concat a b = env_union a b
-
-  let meta ?stmt ?descr n =
-    let data = Meta(stmt,descr) in
-    new_env ~datas:(Node.Map.add n (Bag.elt data) (Node.Map.empty)) ()
-
-  let edge n e =
-    new_env ~succs:(Node.Map.add n e (Node.Map.empty)) ()
-
-  let goto node_orig node_target =
-    edge node_orig (Goto(node_target))
-
-  let branch node_orig predicate node_target_then node_target_else =
-    edge node_orig (Branch(predicate,
-                           Some node_target_then,
-                           Some node_target_else))
+    Node.Map.fold (fun n' preds f ->
+        let l = Node.Map.fold (fun n p acc ->
+            (Conditions.append (have_access n) (!. (Conditions.Have p)))::acc
+          ) preds [] in
+        let f' =
+          Conditions.Branch(access n', either l, Conditions.empty)
+        in
+        let stmt,descr =
+          let bag = match Node.Map.find n' env.datas with
+            | exception Not_found -> Bag.empty
+            | bag -> bag
+          in
+          Bag.fold_left (fun (os,od) b ->
+              match b with
+              | Meta(os',od') ->
+                (if os = None then os' else os),
+                (if od = None then od' else od)
+            ) (None,None) bag in
+        add_cond ?stmt ?descr f f'
+      ) predecessors f
+  in
+
+  (* The transitions *)
+  let to_sequence_basic_forward f =
+    Node.Map.fold (fun n s f ->
+        let node_is_accessible = access n in
+        let f' = match s with
+          | Goto n' | Havoc (_, n') ->
+            (* The havoc is already taken into account during {!domains} *)
+            Conditions.Branch(node_is_accessible,
+                              have_access n',
+                              Conditions.empty)
+          | Branch (c,Some n1,Some n2) ->
+            Conditions.Branch(node_is_accessible,
+                              !. (Conditions.Branch((C.get c),
+                                                    have_access n1,
+                                                    have_access n2)),
+                              Conditions.empty)
+          | Branch(c,Some n1,None) ->
+            Conditions.Branch(node_is_accessible,
+                              Conditions.append
+                                (!. (Conditions.Have (C.get c)))
+                                (have_access n1),
+                              Conditions.empty)
+          | Branch(c,_,Some n1) ->
+            Conditions.Branch(node_is_accessible,
+                              Conditions.append
+                                (!. (Conditions.Have (F.p_not (C.get c))))
+                                (have_access n1),
+                              Conditions.empty)
+          | Branch(_,None,None) -> assert false
+          | Either l ->
+            let l = List.map have_access l in
+            Conditions.Branch(node_is_accessible, either l, Conditions.empty)
+          | Implies l ->
+            let l = List.map
+                (fun (c,n) -> !. (Conditions.Branch (C.get c, have_access n, Conditions.empty)))
+                l in
+            Conditions.Branch(node_is_accessible, Conditions.concat l, Conditions.empty)
+          | Effect (e,n) ->
+            Conditions.Branch(node_is_accessible,
+                              Conditions.append
+                                (!. (Conditions.Have (E.get e)))
+                                (have_access n) ,
+                              Conditions.empty)
+          | Binding (b,n') ->
+            (* For basic: all the variables are important *)
+            let b = !. (Conditions.Have(F.p_conj (Passive.conditions b (fun _ -> true)))) in
+            Conditions.Branch(node_is_accessible,
+                              Conditions.append b (have_access n'),
+                              Conditions.empty) in
+        let stmt,descr =
+          let bag = match Node.Map.find n env.datas with
+            | exception Not_found -> Bag.empty
+            | bag -> bag
+          in
+          Bag.fold_left (fun (os,od) b ->
+              match b with
+              | Meta(os',od') ->
+                (if os = None then os' else os),
+                (if od = None then od' else od)
+            ) (None,None) bag in
+        add_cond ?stmt ?descr f f'
+      ) env.succs f
+  in
+
+  let f = match mode with
+    | `Bool_Backward -> to_sequence_basic_backward f
+    | `Bool_Forward -> to_sequence_basic_forward f
+  in
+  f,Node.Hashtbl.fold Node.Map.add preds Node.Map.empty
+
+module To_tree = struct
+  (* Use a simplified version of "A New Elimination-Based Data Flow Analysis
+      Framework Using Annotated Decomposition Trees" where there is no loop *)
+
+
+  type tree = {
+    c : F.pred (** condition for this tree *) ;
+    q : tree Queue.t (** childrens *) ;
+    mutable fact : F.pred list (** facts at this level *) ;
+  }
 
-  let guard node_orig predicate node_target_then =
-    edge node_orig (Branch(predicate,
-                           Some node_target_then,
-                           None))
+  [@@@ warning "-32"]
+
+  let rec pp_tree
+      ?(pad : (string * string)= ("", ""))
+      (tree : tree) : unit =
+    let pd, pc = pad in
+    Format.printf "%sNode condition: %a @." pd Lang.F.pp_pred tree.c;
+    Format.printf "%sNode fact:%a@."
+      pd
+      (Pretty_utils.pp_list ~sep:"," ~pre:"[" ~suf:"]" Lang.F.pp_pred) tree.fact;
+    let n = Queue.length tree.q - 1 in
+    let _ =
+      Queue.fold (
+        fun i c ->
+          let pad =
+            (pc ^ (if i = n then "`-- " else "|-- "),
+             pc ^ (if i = n then "    " else "|   "))
+          in
+          pp_tree ~pad c;
+          i+1
+      ) 0 tree.q
+    in ()
 
-  let guard' node_orig predicate node_target_else =
-    edge node_orig (Branch(predicate,
-                           None,
-                           Some node_target_else
-                          ))
+  let pp_idoms fmt a =
+    Pretty_utils.pp_array ~sep:";@ " (fun fmt i j -> Format.fprintf fmt "%i -> %i" i j)
+      fmt a
 
-  let either node = function
-    | [] -> nop
-    | [dest] -> goto node dest
-    | node_list -> edge node (Either(node_list))
+  [@@@ warning "+32"]
 
-  let implies node = function
-    | [] -> nop
-    | [g,dest] -> guard node g dest
-    | node_list -> edge node (Implies(node_list))
+  type env_to_sequence_tree = {
+    env: localised_env;
 
-  let memory_effect node1 e node2 =
-    edge node1 (Effect(e, node2))
+    pred: Node.t -> Node.t list;
+    (** predecessors *)
+    topo_order : Node.t -> int;
+    (** topological order *)
+    get_idom_forward: Node.t -> Node.t;
+    (** Immediate dominator forward *)
+    get_idom_backward: int -> int;
+    (** Immediate dominator backward *)
 
-  let assume (predicate:P.t) =
-    if F.is_ptrue (P.get predicate) = Qed.Logic.Yes
-    then nop
-    else new_env ~assumes:(Bag.elt predicate) ()
+    full_conds: Lang.F.pred Node.Hashtbl.t;
+    (** For each node we are going to compute different formulas
+        Necessary conditions of the node from start *)
 
-  let havoc node1 ~memory_effects:node_seq node2 =
-    edge node1 (Havoc(node_seq,node2))
+    conds: Lang.F.pred Node.Hashtbl.t;
+    (** Necessary conditions from its forward idiom *)
 
-  let option_bind ~f = function
-    | None -> None
-    | Some x -> f x
+    subtrees: tree Node.Hashtbl.t;
+    (** To which subtree corresponds this node *)
 
-  let union_opt_or union d1 d2 =
-    match d1, d2 with
-    | Some d1, Some d2 -> Some (union d1 d2)
-    | (Some  _ as d), None | None, (Some _ as d) -> d
-    | None, None -> None
+    root: tree;
+    (** Root the full tree *)
 
-  let union_opt_and union d1 d2 =
-    match d1, d2 with
-    | Some d1, Some d2 -> Some (union d1 d2)
-    | _ -> None
+    eithers: Lang.F.pred Node.Hashtbl.t Node.Hashtbl.t;
+    (** Variable used for the non-deterministic choice of either *)
+  }
 
-  let add_only_if_alive union d1 = function
-    | None -> None
-    | Some d2 -> Some (union d1 d2)
+  let is_after n1 n2 = n1 > n2
+  let is_before n1 n2 = n1 < n2
+
+  let create_env_to_sequence_tree env =
+    (* Compute topological order for immediate dominator computation
+        and the main iteration on nodes
+    *)
+    let node_int,int_node,ordered = topological env in
+    let nb = Node.Hashtbl.length node_int in
+
+    (* We compute the forward immediate dominators (path that use succ)
+        and the backward immediate dominators (path that use pred)
+    *)
+    let predecessors = compute_preds env in
+    let pred n = Node.Hashtbl.find_all predecessors n in
+    let succ n = succs env n in
+    let topo_order = Node.Hashtbl.find node_int in
+    let idoms_forward =
+      idoms ordered topo_order nb ~pred ~is_after in
+    let idoms_backward =
+      idoms (List.rev ordered) topo_order nb ~pred:succ ~is_after:is_before in
+    let get_idom_forward n =
+      Datatype.Int.Hashtbl.find int_node idoms_forward.(topo_order n) in
+    (* Format.printf "@[ordered: %a@]@." (Pretty_utils.pp_list ~sep:"@ " (fun fmt n -> Format.fprintf fmt "%a (%i)" Node.pp n (Node.Hashtbl.find node_int n))) ordered;
+     * Format.printf "@[pred: %a@]@."
+     *   (Pretty_utils.pp_iter2 ~sep:"@ " ~between:":" Node.Hashtbl.iter Node.pp Node.pp) predecessors;
+     * Format.printf "@[idoms forward: @[@[%a@]@]@." _pp_idoms idoms_forward;
+     * Format.printf "@[idoms backward: @[@[%a@]@]@." _pp_idoms idoms_backward; *)
+    {
+      env;
+      pred;
+      topo_order;
+      get_idom_forward;
+      get_idom_backward = (fun i -> idoms_backward.(i));
+      full_conds = Node.Hashtbl.create 10;
+      conds = Node.Hashtbl.create 10;
+      subtrees = Node.Hashtbl.create 10;
+      root = {c = Lang.F.p_true; q = Queue.create (); fact = [] };
+      eithers = Node.Hashtbl.create 10;
+    }, ordered
+
+
+  let either env n last =
+    let h = Node.Hashtbl.memo env.eithers n (fun _ -> Node.Hashtbl.create 10) in
+    Node.Hashtbl.memo h last (fun _ ->
+        let v = F.fresh ~basename:"node"
+            (get_pool ()) Qed.Logic.Bool in
+        F.p_bool (F.e_var v)
+      )
 
-  (** return None when post is not accessible from this node *)
-  let rec effects : type a.  (_,a) env -> node -> node -> S.domain option =
-    fun env post node ->
-    if node = post
-    then Some S.empty
-    else
-      match Node.Map.find node env.succs with
-      | exception Not_found -> None
-      | Goto (node2) ->
-        effects env post node2
-      | Branch (_, node2, node3) ->
-        union_opt_or S.union
-          (option_bind ~f:(effects env post) node2)
-          (option_bind ~f:(effects env post) node3)
-      | Either (l) ->
-        (List.fold_left
-           (fun acc node2 -> union_opt_or S.union
-               acc (effects env post node2))
-           None l)
-      | Implies (l) ->
-        (List.fold_left
-           (fun acc (_,node2) -> union_opt_or S.union
-               acc (effects env post node2))
-           None l)
-      | Effect (memory_effect , node2) ->
-        add_only_if_alive S.union
-          (E.writes memory_effect)
-          (effects env post node2)
-      | Havoc (m, node2) ->
-        union_opt_and S.union
-          (effects env m.post m.pre)
-          (effects env post node2)
-      | Binding (_,node2) ->
-        effects env post node2
-
-  (** restrict a cfg to the nodes accessible from the pre post given,
-      and compute havoc effect *)
-  let restrict (cfg:pre_env) pre posts : restricted_env =
-    let rec walk acc node : restricted_env option =
-      if Node.Map.mem node acc.succs then Some acc
+  (** For a node n *)
+  let iter env n =
+    let idom = env.get_idom_forward n in
+    let rec get_cond acc n' =
+      if n' = idom then acc
       else
-        let new_env edge = new_env ~succs:(Node.Map.add node edge (Node.Map.empty)) () in
-        let r = match Node.Map.find node cfg.succs with
-          | exception Not_found -> None
-          | (Goto (node2) | Effect (_ , node2)) as edge ->
-            union_opt_and env_union
-              (Some (new_env edge))
-              (walk acc node2)
-          | Branch (pred, node2, node3) ->
-            (* it is important to visit all the childrens *)
-            let f acc node =
-              match option_bind ~f:(walk acc) node with
-              | None -> None, acc
-              | Some acc -> node, acc in
-            let node2, acc = f acc node2 in
-            let node3, acc = f acc node3 in
-            if node2 = None && node3 = None then None
-            else Some (env_union acc (new_env (Branch(pred, node2, node3))))
-          | Either (l) ->
-            let acc,l = List.fold_left
-                (fun ((acc,l) as old) node2 ->
-                   match walk acc node2 with
-                   | None -> old
-                   | Some acc -> (acc,node2::l))
-                (acc,[]) l in
-            if l = [] then None
-            else Some (env_union acc (new_env (Either (List.rev l))))
-          | Implies (l) ->
-            let acc,l = List.fold_left
-                (fun ((acc,l) as old) ((_,node2) as e) ->
-                   match walk acc node2 with
-                   | None -> old
-                   | Some acc -> (acc,e::l))
-                (acc,[]) l in
-            if l = [] then None
-            else Some (env_union acc (new_env (Implies (List.rev l))))
-          | Havoc (m, node2) ->
-            match effects cfg m.post m.pre with
-            | None -> None
-            | Some eff ->
-              union_opt_and env_union
-                (Some (new_env (Havoc(eff,node2))))
-                (walk acc node2)
-        in
-        if Node.Set.mem node posts && r = None
-        then Some acc
-        else r
-    in
-    match walk (new_env ()) pre with
-    | None -> (new_env ())
-    | Some acc ->
-      { succs = acc.succs;
-        datas = Node.Map.inter (fun _ _ v -> v) acc.succs cfg.datas;
-        assumes = Bag.filter (fun (seq,_) ->
-            Node.Map.subset (fun _ _ _ -> true)
-              seq acc.succs) cfg.assumes;
-        tmpnodes = cfg.tmpnodes;
-      }
-
-  (** succ is decreasing for this order *)
-  let topological (type a) (type b) (cfg:(a,b) env) =
-    let module G = struct
-      type t = (a,b) env
-      module V = struct let hash = Hashtbl.hash include Node end
-      let iter_vertex f cfg =
-        let h = Node.Hashtbl.create 10 in
-        let replace n = Node.Hashtbl.replace h n () in
-        Node.Map.iter (fun k _ -> replace k; iter_succs_e replace cfg k) cfg.succs;
-        Node.Hashtbl.iter (fun k () -> f k) h
-      let iter_succ = iter_succs_e
-    end in
-    let module T = Graph.Topological.Make(G) in
-    let h  = Node.Hashtbl.create 10 in
-    let h' = Datatype.Int.Hashtbl.create 10 in
-    let c = ref (-1) in
-    let l = ref [] in
-    T.iter (fun n -> l := n::!l; incr c; Node.Hashtbl.add h n !c; Datatype.Int.Hashtbl.add h' !c n) cfg;
-    h,h',List.rev !l
-
-  (** topo_list: elements in topological order
-      topo_order: post-order mapping
-      nb: number of elements *)
-  let idoms topo_list topo_order nb ~pred ~is_after =
-    let a = Array.make nb (-1) in
-    let iter n =
-      let first,preds = match pred n with
-        | [] -> topo_order n, []
-        | f::p -> topo_order f, List.map topo_order p
-      in
-      let rec find_common n1 n2 =
-        if n1 = n2 then n1
-        else if is_after n1 n2 then find_common a.(n1) n2
-        else find_common n1 a.(n2) in
-      let idom = List.fold_left find_common first preds in
-      a.(topo_order n) <- idom
-    in
-    List.iter iter topo_list;
-    a
-
-  let find_def ~def x t = try Node.Map.find x t with Not_found -> def
-
-  let rec remove_dumb_gotos (env:restricted_env) : Node.t Node.Map.t * restricted_env =
-    let add_map m acc = Node.Map.fold (fun n _ acc -> Node.Set.add n acc) m acc in
-    let used_nodes =
-      Bag.fold_left (fun acc p -> add_map (P.reads p) acc) Node.Set.empty env.assumes
+        let acc = F.p_and acc (Node.Hashtbl.find env.conds n') in
+        get_cond acc (env.get_idom_forward n')
     in
-    let used_nodes = add_map env.datas used_nodes in
-    let how_many_preds = Node.Hashtbl.create 10 in
-    let incr_how_many_preds n =
-      Node.Hashtbl.replace how_many_preds n (succ (Node.Hashtbl.find_def how_many_preds n 0))
-    in
-    let subst =
-      Node.Map.fold (fun n e acc ->
-          iter_succs incr_how_many_preds e;
-          match (e:(_,without_bindings) edge) with
-          | Goto n' when not (Node.Set.mem n used_nodes) ->
-            Node.Map.add n n' acc
-          | Goto _
-          | Branch (_,_,_)
-          | Either _
-          | Implies _
-          | Effect (_,_)
-          | Havoc (_,_) -> acc)
-        env.succs Node.Map.empty
-    in
-    let subst =
-      let rec compress n =
-        match Node.Map.find n subst with
-        | exception Not_found -> n
-        | n -> compress n
-      in
-      Node.Map.map (fun _ n' -> compress n') subst
-    in
-    let find n = find_def ~def:n n subst in
-    (* detect either that could be transformed in branch *)
-    let to_remove = Node.Hashtbl.create 10 in
-    Node.Map.iter (fun _ e ->
-        match (e:(_,without_bindings) edge) with
-        | Either [a;b] when Node.Hashtbl.find how_many_preds a = 1 &&
-                            Node.Hashtbl.find how_many_preds b = 1 &&
-                            not (Node.Set.mem a used_nodes) &&
-                            not (Node.Set.mem b used_nodes) &&
-                            Node.Set.mem a env.tmpnodes &&
-                            Node.Set.mem b env.tmpnodes &&
-                            true
-          ->
-          begin
-            let find_opt k m =
-              match Node.Map.find k m with
-              | exception Not_found -> None
-              | v -> Some v
-            in
-            match find_opt a env.succs, find_opt b env.succs with
-            | Some Branch(c,Some n1, None), Some Branch(c',None, Some n2)
-            | Some Branch(c,None, Some n2), Some Branch(c',Some n1,None) when C.equal c c' ->
-              let n1 = find n1 in
-              let n2 = find n2 in
-              let br = Branch(c,Some n1, Some n2) in
-              Node.Hashtbl.add to_remove a br;
-              Node.Hashtbl.add to_remove b br
-            | _ -> ()
-          end
-        | Goto _
-        | Branch (_,_,_)
-        | Effect (_,_)
-        | Either _
-        | Implies _
-        | Havoc (_,_) -> ()
-      ) env.succs;
-    (* substitute and remove *)
-    let succs = Node.Map.mapq (fun n e ->
-        match (e:(_,without_bindings) edge) with
-        | _ when Node.Hashtbl.mem to_remove n -> None
-        | Goto _ when not (Node.Set.mem n used_nodes) -> None
-        | Goto n' ->
-          let n'' = find n' in
-          if Node.equal n' n'' then Some e
-          else Some (Goto n'')
-        | Branch (c,n1,n2) ->
-          let n1' = Option.map find n1 in
-          let n2' = Option.map find n2 in
-          if Option.equal Node.equal n1 n1' && Option.equal Node.equal n2 n2'
-          then Some e
-          else Some (Branch(c,n1',n2'))
-        | Either l ->
-          let l' = List.map find l in
-          let l' = List.sort_uniq Node.compare l' in
-          begin match l' with
-            | [] -> assert false (* absurd: Either after restricted has at least one successor *)
-            | [a] -> Some (Goto a)
-            | [a;_] when Node.Hashtbl.mem to_remove a ->
-              let br = Node.Hashtbl.find to_remove a in
-              Some br
-            | l' -> Some (Either l')
-          end
-        | Implies l ->
-          let l' = List.map (fun (g,n) -> (g,find n)) l in
-          Some (Implies l')
-        | Effect (ef,n') ->
-          let n'' = find n' in
-          if Node.equal n' n'' then Some e
-          else Some (Effect(ef,n''))
-        | Havoc (h,n') ->
-          let n'' = find n' in
-          if Node.equal n' n'' then Some e
-          else Some (Havoc(h,n''))
-      )
-        env.succs
-    in
-    let env = {env with succs} in
-    if Node.Map.is_empty subst
-    then subst, env
-    else
-      let subst', env = remove_dumb_gotos env in
-      let subst = Node.Map.map (fun _ n' -> find_def ~def:n' n' subst') subst in
-      Node.Map.merge (fun _ a b ->
-          match a, b with
-          | Some _, Some _ -> assert false (* the elements are remove in the new env *)
-          | Some x, None | None, Some x -> Some x
-          | None, None -> assert false
-        ) subst subst', env
-
-  let allocate domain sigma =
-    S.Chunk.Set.iter (fun chunk -> ignore (S.get sigma chunk)) domain
-
-  let domains (env : restricted_env) reads pre : localised_env * S.t Node.Map.t =
-    let visited = ref Node.Map.empty in
-    let new_succs = ref Node.Map.empty in
-    let add_edge node edge = new_succs := Node.Map.add node edge !new_succs in
-    let add_binding_edge n (p: Passive.t) =
-      if Passive.is_empty p then n
+    (* find all the conditions that keep the path toward n, i.e.
+        the condition of the nodes that are not dominated backwardly
+        (for which not all the nodes goes to n)
+    *)
+    let rec find_frontiere last acc n' =
+      if (env.topo_order n) <=  env.get_idom_backward (env.topo_order n')
+      then
+        let cond = get_cond F.p_true n' in
+        let branch = match Node.Map.find n' env.env.succs with
+          | exception Not_found -> F.p_true
+          | Goto _ | Havoc (_, _) -> F.p_true
+          | Branch (c,Some n'',Some _) when Node.equal n'' last -> C.get c
+          | Branch (c,Some _,Some n'') when Node.equal n'' last -> F.p_not (C.get c)
+          | Branch (_,_,_) -> F.p_true
+          | Either _ -> either env n' last
+          | Implies l -> List.fold_left (fun acc (c,n) ->
+              if n = last then C.get c
+              else acc) F.p_true l
+          | Effect (_,_) -> F.p_true
+          | Binding (_,_) -> F.p_true
+        in
+        F.p_or acc (F.p_and branch cond)
       else
-        let n' = Node.node_internal () in
-        add_edge n' (Binding(p,n));
-        n'
+        List.fold_left (find_frontiere n') acc (env.pred n')
     in
-    let rec aux node : S.t =
-      try Node.Map.find node !visited
-      with Not_found ->
-        let dom = find_def ~def:S.empty node reads in
-        let ret =
-          match Node.Map.find node env.succs with
-          | exception Not_found ->
-            (* posts node *)
-            let s1 = S.create () in
-            allocate dom s1;
-            s1
-          | Goto (node2) ->
-            let s1 = S.copy (aux node2) in
-            allocate dom s1;
-            add_edge node (Goto node2);
-            s1
-          | Branch (pred, node2, node3) ->
-            let dom = (S.union (C.reads pred) dom) in
-            begin match node2, node3 with
-              | (None, Some next) | (Some next, None) ->
-                let s1 = S.copy (aux next) in
-                allocate dom s1;
-                let pred = C.relocate s1 pred in
-                add_edge node (Branch(pred,node2,node3));
-                s1
-              | Some node2, Some node3 ->
-                let s2 = aux node2 in
-                let s3 = aux node3 in
-                let s1,p2,p3 = S.merge s2 s3 in
-                allocate dom s1;
-                let node2' = add_binding_edge node2 p2 in
-                let node3' = add_binding_edge node3 p3 in
-                let pred = C.relocate s1 pred in
-                add_edge node (Branch(pred,Some node2',Some node3'));
-                s1
-              | _ -> assert false
-            end
-          | Either (l) ->
-            let s1, pl = S.merge_list (List.map aux l) in
-            allocate dom s1;
-            let l = List.map2 add_binding_edge l pl in
-            add_edge node (Either l);
-            s1
-          | Implies (l) ->
-            let dom =
-              List.fold_left (fun acc (c,_) -> S.union (C.reads c) acc) dom l
-            in
-            let s1, pl = S.merge_list (List.map (fun (_,n) -> aux n) l) in
-            allocate dom s1;
-            let l = List.map2 (fun (c,a) b ->
-                let a = add_binding_edge a b in
-                let c = C.relocate s1 c in
-                (c,a)) l pl
-            in
-            add_edge node (Implies l);
-            s1
-          | Effect (memory_effect , node2) ->
-            let s2 = aux node2 in
-            let s1 = S.remove_chunks s2 (E.writes memory_effect) in
-            allocate dom s1;
-            allocate (E.reads memory_effect) s1;
-            let memory_effect = E.relocate {pre=s1;post=s2} memory_effect in
-            add_edge node (Effect(memory_effect,node2));
-            s1
-          | Havoc (eff, node2) ->
-            let s2 = aux node2 in
-            let s1 = S.havoc s2 eff in
-            allocate dom s1;
-            add_edge node (Havoc(eff,node2));
-            s1
-        in
-        visited := Node.Map.add node ret !visited;
-        ret
+    let c, q =
+      if Node.equal idom n
+      then
+        (* it is the root *)
+        begin
+          Node.Hashtbl.add env.full_conds n F.p_true;
+          Node.Hashtbl.add env.conds n F.p_true;
+          F.p_true, env.root.q
+        end
+      else
+        let c = List.fold_left (find_frontiere n) F.p_false (env.pred n) in
+        (* Format.printf "for %a c=%a@." Node.pp n Lang.F.pp_pred c; *)
+        Node.Hashtbl.add env.conds n c;
+        Node.Hashtbl.add env.full_conds n (F.p_and c (Node.Hashtbl.find env.full_conds idom));
+        let p = Node.Hashtbl.find env.subtrees idom in
+        c,p.q
     in
-    ignore (aux pre);
-    let sigmas = !visited in
-    let new_env =
-      {succs = !new_succs;
-       datas = env.datas;
-       assumes = Bag.map (fun e -> P.relocate sigmas e) env.assumes;
-       tmpnodes = env.tmpnodes;
-      } in
-    new_env, sigmas
-
-  let compute_preds env =
-    let h = Node.Hashtbl.create 10 in
-    let add = Node.Hashtbl.add h in
-    Node.Map.iter (fun n s ->
-        match s with
-        | Goto n1 | Havoc (_, n1) | Effect (_,n1) | Binding (_,n1) -> add n1 n
-        | Branch (_,Some n1,Some n2) ->
-          add n1 n;
-          add n2 n
-        | Branch(_,Some n1,None) -> add n1 n
-        | Branch(_,None,Some n1) -> add n1 n
-        | Branch(_,None,None) -> ()
-        | Either l -> List.iter (fun n1 -> add n1 n) l
-        | Implies l -> List.iter (fun (_,n1) -> add n1 n) l
-      ) env.succs;
-    h
-
-  let to_sequence_bool ~mode pre posts env : Conditions.sequence * F.pred Node.Map.t =
-    let preds = Node.Hashtbl.create 10 in
-    let access n = Node.Hashtbl.memo preds n
-        (fun _ ->
-           let v = F.fresh ~basename:"node"
-               (get_pool ()) Qed.Logic.Bool in
-           F.p_bool (F.e_var v))
+    let fact = match Node.Map.find n env.env.succs with
+      | exception Not_found -> F.p_true
+      | Goto _ | Havoc (_, _) -> F.p_true
+      | Branch (c,Some _,None) -> C.get c
+      | Branch (c,None,Some _) -> F.p_not (C.get c)
+      | Branch (_,_,_) -> F.p_true
+      | Either _ -> F.p_true
+      | Implies _ -> F.p_true
+      | Effect (e,_) -> E.get e
+      | Binding (b,_) -> F.p_conj (Passive.conditions b (fun _ -> true))
     in
-    let (!.) c = (Conditions.sequence [Conditions.step c]) in
-    let have_access n = !. (Conditions.Have (access n)) in
-    let add_cond ?descr ?stmt f cond =
+    (* Format.printf "Here: For %a idoms=%a c=%a fact=%a@." Node.pp n Node.pp idom F.pp_pred c F.pp_pred fact; *)
+    let t = {c = c; q = Queue.create (); fact = [fact]} in
+    Queue.push t q;
+    Node.Hashtbl.add env.subtrees n t
+
+  let add_cond ?descr ?stmt f cond =
+    match cond with
+    | Conditions.Have c when F.is_ptrue c = Qed.Logic.Yes -> f
+    | _ ->
       Conditions.append (Conditions.sequence [Conditions.step ?descr ?stmt cond]) f
-    in
-    let either = function
-      | [] -> !. (Conditions.Have F.p_false)
-      | [a] -> a
-      | l -> !. (Conditions.Either l)
-    in
+
+  let access env n = Node.Hashtbl.find env.full_conds n
+
+  let get_latest_node env  = function
+    | [] -> env.root
+    | a::l ->
+      let n = List.fold_left (fun a e ->
+          if env.topo_order a < env.topo_order e then e else a
+        ) a l in
+      Node.Hashtbl.find env.subtrees n
+
+  (** Add each assume to the sub-tree corresponding to the latest
+      node it uses. The assumes are true if all their nodes are
+      accessible *)
+  let add_assumes_fact env = Bag.iter (fun p ->
+      let nodes = P.nodes_list p in
+      let nodes_are_accessible =
+        (* TODO: don't add the condition of access of the node that are dominators of latest *)
+        List.fold_left (fun acc n -> F.p_and (access env n) acc) F.p_true nodes in
+      let f' = F.p_imply nodes_are_accessible (P.get p) in
+      let t = get_latest_node env nodes in
+      t.fact <- f' :: t.fact
+    ) env.env.assumes
+
+  (** convert the tree to formula *)
+  let rec convert t f =
+    let f' =
+      if t.fact = []
+      then Conditions.empty
+      else List.fold_left (fun acc e -> add_cond acc (Conditions.Have e)) Conditions.empty t.fact in
+    let f' = Queue.fold (fun f' t -> convert t f') f' t.q in
+    match F.is_ptrue t.c with
+    | Qed.Logic.Yes -> Conditions.concat [f;f']
+    | Qed.Logic.No -> f
+    | Qed.Logic.Maybe ->
+      add_cond f (Conditions.Branch(t.c,f',Conditions.empty))
+
+  let to_sequence_tree _ posts env =
+    let env,ordered = create_env_to_sequence_tree env in
+    (* Iterate in topo order the vertex.
+        Except for root, the tree of the vertex is the one of its immediate dominator forward.
+    *)
+    List.iter (iter env) ordered;
     let f = Conditions.empty in
-    (* The start state is accessible *)
-    let pre = Conditions.Have (access pre) in
-    let f = add_cond f pre in
     (* The posts state are accessible *)
     let f = Node.Set.fold
-        (fun n f -> add_cond f (Conditions.Have (access n)))
+        (fun n f -> add_cond f (Conditions.Have (access env n)))
         posts f in
-    (* The assumes are true if all their nodes are accessible *)
-    let f =
-      Bag.fold_left (fun f p ->
-          let nodes_are_accessible =
-            Node.Map.fold (fun n _ acc -> F.p_and (access n) acc)
-              (P.reads p) F.p_true in
-          let f' = F.p_imply nodes_are_accessible (P.get p) in
-          add_cond f (Conditions.Have f')
-        ) f env.assumes in
-
-    (* compute predecessors *)
-    let to_sequence_basic_backward f =
-      let predecessors = Node.Map.fold (fun n s acc ->
-          let add acc n' p =
-            Node.Map.change (fun _ (n,p) -> function
-                | None -> Some (Node.Map.add n p Node.Map.empty)
-                | Some s -> Some (Node.Map.add n p s)) n' (n,p) acc
-          in
-          match s with
-          | Goto n' | Havoc (_, n') -> add acc n' F.p_true
-          | Branch (c,Some n1,Some n2) ->
-            let c = P.get c in
-            add (add acc n1 c) n2 (F.p_not c)
-          | Branch(c,Some n1,None) -> add acc n1 (P.get c)
-          | Branch(c,None,Some n1) -> add acc n1 (F.p_not (P.get c))
-          | Branch(_,None,None) -> acc
-          | Either l -> List.fold_left (fun acc e -> add acc e F.p_true) acc l
-          | Implies l -> List.fold_left (fun acc (c,e) -> add acc e (P.get c)) acc l
-          | Effect (e,n') -> add acc n' (E.get e)
-          | Binding (b,n') ->
-            let b = F.p_conj (Passive.conditions b (fun _ -> true)) in
-            add acc n' b
-        ) env.succs Node.Map.empty
-      in
-      Node.Map.fold (fun n' preds f ->
-          let l = Node.Map.fold (fun n p acc ->
-              (Conditions.append (have_access n) (!. (Conditions.Have p)))::acc
-            ) preds [] in
-          let f' =
-            Conditions.Branch(access n', either l, Conditions.empty)
-          in
-          let stmt,descr =
-            let bag = match Node.Map.find n' env.datas with
-              | exception Not_found -> Bag.empty
-              | bag -> bag
-            in
-            Bag.fold_left (fun (os,od) b ->
-                match b with
-                | Meta(os',od') ->
-                  (if os = None then os' else os),
-                  (if od = None then od' else od)
-              ) (None,None) bag in
-          add_cond ?stmt ?descr f f'
-        ) predecessors f
-    in
+    (* For all either one of the condition is true *)
+    let f = Node.Hashtbl.fold (fun _ h f ->
+        let p = Node.Hashtbl.fold (fun _ t p -> F.p_or p t) h F.p_false in
+        add_cond f (Conditions.Have p)
+      ) env.eithers f in
+    add_assumes_fact env;
+    let f = convert env.root f in
+    f, Node.Hashtbl.fold Node.Map.add env.full_conds Node.Map.empty
+end
 
-    (* The transitions *)
-    let to_sequence_basic_forward f =
-      Node.Map.fold (fun n s f ->
-          let node_is_accessible = access n in
-          let f' = match s with
-            | Goto n' | Havoc (_, n') ->
-              (* The havoc is already taken into account during {!domains} *)
-              Conditions.Branch(node_is_accessible,
-                                have_access n',
-                                Conditions.empty)
-            | Branch (c,Some n1,Some n2) ->
-              Conditions.Branch(node_is_accessible,
-                                !. (Conditions.Branch((C.get c),
-                                                      have_access n1,
-                                                      have_access n2)),
-                                Conditions.empty)
-            | Branch(c,Some n1,None) ->
-              Conditions.Branch(node_is_accessible,
-                                Conditions.append
-                                  (!. (Conditions.Have (C.get c)))
-                                  (have_access n1),
-                                Conditions.empty)
-            | Branch(c,_,Some n1) ->
-              Conditions.Branch(node_is_accessible,
-                                Conditions.append
-                                  (!. (Conditions.Have (F.p_not (C.get c))))
-                                  (have_access n1),
-                                Conditions.empty)
-            | Branch(_,None,None) -> assert false
-            | Either l ->
-              let l = List.map have_access l in
-              Conditions.Branch(node_is_accessible, either l, Conditions.empty)
-            | Implies l ->
-              let l = List.map
-                  (fun (c,n) -> !. (Conditions.Branch (C.get c, have_access n, Conditions.empty)))
-                  l in
-              Conditions.Branch(node_is_accessible, Conditions.concat l, Conditions.empty)
-            | Effect (e,n) ->
-              Conditions.Branch(node_is_accessible,
-                                Conditions.append
-                                  (!. (Conditions.Have (E.get e)))
-                                  (have_access n) ,
-                                Conditions.empty)
-            | Binding (b,n') ->
-              (* For basic: all the variables are important *)
-              let b = !. (Conditions.Have(F.p_conj (Passive.conditions b (fun _ -> true)))) in
-              Conditions.Branch(node_is_accessible,
-                                Conditions.append b (have_access n'),
-                                Conditions.empty) in
-          let stmt,descr =
-            let bag = match Node.Map.find n env.datas with
-              | exception Not_found -> Bag.empty
-              | bag -> bag
-            in
-            Bag.fold_left (fun (os,od) b ->
-                match b with
-                | Meta(os',od') ->
-                  (if os = None then os' else os),
-                  (if od = None then od' else od)
-              ) (None,None) bag in
-          add_cond ?stmt ?descr f f'
-        ) env.succs f
+let compile : ?name:string -> ?mode:mode -> node -> Node.Set.t -> domain Node.Map.t ->
+  cfg -> F.pred Node.Map.t * sigma Node.Map.t * Conditions.sequence =
+  fun ?(name="cfg") ?(mode=`Bool_Forward) pre posts user_reads env ->
+  if Wp_parameters.has_dkey dkey then
+    Format.printf "@[0) pre:%a post:%a@]@."
+      Node.pp pre (Pretty_utils.pp_iter ~sep:"@ " Node.Set.iter Node.pp) posts;
+  if Wp_parameters.has_dkey dkey then
+    Format.printf "@[1) %a@]@." pretty_env env;
+  (* restrict environment to useful node and compute havoc effects *)
+  let env = restrict env pre posts in
+  if Wp_parameters.has_dkey dkey then
+    Format.printf "@[2) %a@]@." pretty_env env;
+  if Node.Map.is_empty env.succs then
+    Node.Map.empty,Node.Map.empty,
+    Conditions.sequence [Conditions.step (Conditions.Have(F.p_false))]
+  else
+    (* Simplify *)
+    let subst,env =
+      if true
+      then remove_dumb_gotos env
+      else Node.Map.empty, env
     in
-
-    let f = match mode with
-      | `Bool_Backward -> to_sequence_basic_backward f
-      | `Bool_Forward -> to_sequence_basic_forward f
+    let pre = find_def ~def:pre pre subst in
+    (* Substitute in user_reads *)
+    let user_reads =
+      Node.Map.fold
+        (fun n n' acc ->
+           match Node.Map.find n user_reads with
+           | exception Not_found -> acc
+           | domain ->
+             let domain' =
+               try
+                 Sigma.union (Node.Map.find n' acc) domain
+               with Not_found -> domain
+             in
+             Node.Map.add n' domain' acc)
+        subst user_reads
     in
-    f,Node.Hashtbl.fold Node.Map.add preds Node.Map.empty
-
-  module To_tree = struct
-    (* Use a simplified version of "A New Elimination-Based Data Flow Analysis
-        Framework Using Annotated Decomposition Trees" where there is no loop *)
-
-
-    type tree = {
-      c : F.pred (** condition for this tree *) ;
-      q : tree Queue.t (** childrens *) ;
-      mutable fact : F.pred list (** facts at this level *) ;
-    }
-
-    [@@@ warning "-32"]
-
-    let rec pp_tree
-        ?(pad : (string * string)= ("", ""))
-        (tree : tree) : unit =
-      let pd, pc = pad in
-      Format.printf "%sNode condition: %a @." pd Lang.F.pp_pred tree.c;
-      Format.printf "%sNode fact:%a@."
-        pd
-        (Pretty_utils.pp_list ~sep:"," ~pre:"[" ~suf:"]" Lang.F.pp_pred) tree.fact;
-      let n = Queue.length tree.q - 1 in
-      let _ =
-        Queue.fold (
-          fun i c ->
-            let pad =
-              (pc ^ (if i = n then "`-- " else "|-- "),
-               pc ^ (if i = n then "    " else "|   "))
-            in
-            pp_tree ~pad c;
-            i+1
-        ) 0 tree.q
-      in ()
-
-    let pp_idoms fmt a =
-      Pretty_utils.pp_array ~sep:";@ " (fun fmt i j -> Format.fprintf fmt "%i -> %i" i j)
-        fmt a
-
-    [@@@ warning "+32"]
-
-    type env_to_sequence_tree = {
-      env: localised_env;
-
-      pred: Node.t -> Node.t list;
-      (** predecessors *)
-      topo_order : Node.t -> int;
-      (** topological order *)
-      get_idom_forward: Node.t -> Node.t;
-      (** Immediate dominator forward *)
-      get_idom_backward: int -> int;
-      (** Immediate dominator backward *)
-
-      full_conds: Lang.F.pred Node.Hashtbl.t;
-      (** For each node we are going to compute different formulas
-          Necessary conditions of the node from start *)
-
-      conds: Lang.F.pred Node.Hashtbl.t;
-      (** Necessary conditions from its forward idiom *)
-
-      subtrees: tree Node.Hashtbl.t;
-      (** To which subtree corresponds this node *)
-
-      root: tree;
-      (** Root the full tree *)
-
-      eithers: Lang.F.pred Node.Hashtbl.t Node.Hashtbl.t;
-      (** Variable used for the non-deterministic choice of either *)
-    }
-
-    let is_after n1 n2 = n1 > n2
-    let is_before n1 n2 = n1 < n2
-
-    let create_env_to_sequence_tree env =
-      (* Compute topological order for immediate dominator computation
-          and the main iteration on nodes
-      *)
-      let node_int,int_node,ordered = topological env in
-      let nb = Node.Hashtbl.length node_int in
-
-      (* We compute the forward immediate dominators (path that use succ)
-          and the backward immediate dominators (path that use pred)
-      *)
-      let predecessors = compute_preds env in
-      let pred n = Node.Hashtbl.find_all predecessors n in
-      let succ n = succs env n in
-      let topo_order = Node.Hashtbl.find node_int in
-      let idoms_forward =
-        idoms ordered topo_order nb ~pred ~is_after in
-      let idoms_backward =
-        idoms (List.rev ordered) topo_order nb ~pred:succ ~is_after:is_before in
-      let get_idom_forward n =
-        Datatype.Int.Hashtbl.find int_node idoms_forward.(topo_order n) in
-      (* Format.printf "@[ordered: %a@]@." (Pretty_utils.pp_list ~sep:"@ " (fun fmt n -> Format.fprintf fmt "%a (%i)" Node.pp n (Node.Hashtbl.find node_int n))) ordered;
-       * Format.printf "@[pred: %a@]@."
-       *   (Pretty_utils.pp_iter2 ~sep:"@ " ~between:":" Node.Hashtbl.iter Node.pp Node.pp) predecessors;
-       * Format.printf "@[idoms forward: @[@[%a@]@]@." _pp_idoms idoms_forward;
-       * Format.printf "@[idoms backward: @[@[%a@]@]@." _pp_idoms idoms_backward; *)
-      {
-        env;
-        pred;
-        topo_order;
-        get_idom_forward;
-        get_idom_backward = (fun i -> idoms_backward.(i));
-        full_conds = Node.Hashtbl.create 10;
-        conds = Node.Hashtbl.create 10;
-        subtrees = Node.Hashtbl.create 10;
-        root = {c = Lang.F.p_true; q = Queue.create (); fact = [] };
-        eithers = Node.Hashtbl.create 10;
-      }, ordered
-
-
-    let either env n last =
-      let h = Node.Hashtbl.memo env.eithers n (fun _ -> Node.Hashtbl.create 10) in
-      Node.Hashtbl.memo h last (fun _ ->
-          let v = F.fresh ~basename:"node"
-              (get_pool ()) Qed.Logic.Bool in
-          F.p_bool (F.e_var v)
-        )
-
-    (** For a node n *)
-    let iter env n =
-      let idom = env.get_idom_forward n in
-      let rec get_cond acc n' =
-        if n' = idom then acc
-        else
-          let acc = F.p_and acc (Node.Hashtbl.find env.conds n') in
-          get_cond acc (env.get_idom_forward n')
-      in
-      (* find all the conditions that keep the path toward n, i.e.
-          the condition of the nodes that are not dominated backwardly
-          (for which not all the nodes goes to n)
-      *)
-      let rec find_frontiere last acc n' =
-        if (env.topo_order n) <=  env.get_idom_backward (env.topo_order n')
-        then
-          let cond = get_cond F.p_true n' in
-          let branch = match Node.Map.find n' env.env.succs with
-            | exception Not_found -> F.p_true
-            | Goto _ | Havoc (_, _) -> F.p_true
-            | Branch (c,Some n'',Some _) when Node.equal n'' last -> C.get c
-            | Branch (c,Some _,Some n'') when Node.equal n'' last -> F.p_not (C.get c)
-            | Branch (_,_,_) -> F.p_true
-            | Either _ -> either env n' last
-            | Implies l -> List.fold_left (fun acc (c,n) ->
-                if n = last then C.get c
-                else acc) F.p_true l
-            | Effect (_,_) -> F.p_true
-            | Binding (_,_) -> F.p_true
-          in
-          F.p_or acc (F.p_and branch cond)
-        else
-          List.fold_left (find_frontiere n') acc (env.pred n')
-      in
-      let c, q =
-        if Node.equal idom n
-        then
-          (* it is the root *)
-          begin
-            Node.Hashtbl.add env.full_conds n F.p_true;
-            Node.Hashtbl.add env.conds n F.p_true;
-            F.p_true, env.root.q
-          end
-        else
-          let c = List.fold_left (find_frontiere n) F.p_false (env.pred n) in
-          (* Format.printf "for %a c=%a@." Node.pp n Lang.F.pp_pred c; *)
-          Node.Hashtbl.add env.conds n c;
-          Node.Hashtbl.add env.full_conds n (F.p_and c (Node.Hashtbl.find env.full_conds idom));
-          let p = Node.Hashtbl.find env.subtrees idom in
-          c,p.q
-      in
-      let fact = match Node.Map.find n env.env.succs with
-        | exception Not_found -> F.p_true
-        | Goto _ | Havoc (_, _) -> F.p_true
-        | Branch (c,Some _,None) -> C.get c
-        | Branch (c,None,Some _) -> F.p_not (C.get c)
-        | Branch (_,_,_) -> F.p_true
-        | Either _ -> F.p_true
-        | Implies _ -> F.p_true
-        | Effect (e,_) -> E.get e
-        | Binding (b,_) -> F.p_conj (Passive.conditions b (fun _ -> true))
-      in
-      (* Format.printf "Here: For %a idoms=%a c=%a fact=%a@." Node.pp n Node.pp idom F.pp_pred c F.pp_pred fact; *)
-      let t = {c = c; q = Queue.create (); fact = [fact]} in
-      Queue.push t q;
-      Node.Hashtbl.add env.subtrees n t
-
-    let add_cond ?descr ?stmt f cond =
-      match cond with
-      | Conditions.Have c when F.is_ptrue c = Qed.Logic.Yes -> f
-      | _ ->
-        Conditions.append (Conditions.sequence [Conditions.step ?descr ?stmt cond]) f
-
-    let access env n = Node.Hashtbl.find env.full_conds n
-
-    let get_latest_node env  = function
-      | [] -> env.root
-      | a::l ->
-        let n = List.fold_left (fun a e ->
-            if env.topo_order a < env.topo_order e then e else a
-          ) a l in
-        Node.Hashtbl.find env.subtrees n
-
-    (** Add each assume to the sub-tree corresponding to the latest
-        node it uses. The assumes are true if all their nodes are
-        accessible *)
-    let add_assumes_fact env = Bag.iter (fun p ->
-        let nodes = P.nodes_list p in
-        let nodes_are_accessible =
-          (* TODO: don't add the condition of access of the node that are dominators of latest *)
-          List.fold_left (fun acc n -> F.p_and (access env n) acc) F.p_true nodes in
-        let f' = F.p_imply nodes_are_accessible (P.get p) in
-        let t = get_latest_node env nodes in
-        t.fact <- f' :: t.fact
-      ) env.env.assumes
-
-    (** convert the tree to formula *)
-    let rec convert t f =
-      let f' =
-        if t.fact = []
-        then Conditions.empty
-        else List.fold_left (fun acc e -> add_cond acc (Conditions.Have e)) Conditions.empty t.fact in
-      let f' = Queue.fold (fun f' t -> convert t f') f' t.q in
-      match F.is_ptrue t.c with
-      | Qed.Logic.Yes -> Conditions.concat [f;f']
-      | Qed.Logic.No -> f
-      | Qed.Logic.Maybe ->
-        add_cond f (Conditions.Branch(t.c,f',Conditions.empty))
-
-    let to_sequence_tree _ posts env =
-      let env,ordered = create_env_to_sequence_tree env in
-      (* Iterate in topo order the vertex.
-          Except for root, the tree of the vertex is the one of its immediate dominator forward.
-      *)
-      List.iter (iter env) ordered;
-      let f = Conditions.empty in
-      (* The posts state are accessible *)
-      let f = Node.Set.fold
-          (fun n f -> add_cond f (Conditions.Have (access env n)))
-          posts f in
-      (* For all either one of the condition is true *)
-      let f = Node.Hashtbl.fold (fun _ h f ->
-          let p = Node.Hashtbl.fold (fun _ t p -> F.p_or p t) h F.p_false in
-          add_cond f (Conditions.Have p)
-        ) env.eithers f in
-      add_assumes_fact env;
-      let f = convert env.root f in
-      f, Node.Hashtbl.fold Node.Map.add env.full_conds Node.Map.empty
-  end
-
-  let compile : ?name:string -> ?mode:mode -> node -> Node.Set.t -> S.domain Node.Map.t ->
-    cfg -> F.pred Node.Map.t * S.t Node.Map.t * Conditions.sequence =
-    fun ?(name="cfg") ?(mode=`Bool_Forward) pre posts user_reads env ->
-    if Wp_parameters.has_dkey dkey then
-      Format.printf "@[0) pre:%a post:%a@]@."
-        Node.pp pre (Pretty_utils.pp_iter ~sep:"@ " Node.Set.iter Node.pp) posts;
-    if Wp_parameters.has_dkey dkey then
-      Format.printf "@[1) %a@]@." pretty_env env;
-    (* restrict environment to useful node and compute havoc effects *)
-    let env = restrict env pre posts in
+    (* For each node what must be read for assumes *)
+    let reads =
+      Bag.fold_left (fun acc e ->
+          Node.Map.union
+            (fun _ -> Sigma.union) acc
+            (P.reads e))
+        user_reads env.assumes in
+    (* compute sigmas and relocate them *)
+    let env, sigmas = domains env reads pre in
     if Wp_parameters.has_dkey dkey then
-      Format.printf "@[2) %a@]@." pretty_env env;
-    if Node.Map.is_empty env.succs then
-      Node.Map.empty,Node.Map.empty,
-      Conditions.sequence [Conditions.step (Conditions.Have(F.p_false))]
-    else
-      (* Simplify *)
-      let subst,env =
-        if true
-        then remove_dumb_gotos env
-        else Node.Map.empty, env
-      in
-      let pre = find_def ~def:pre pre subst in
-      (* Substitute in user_reads *)
-      let user_reads =
-        Node.Map.fold
-          (fun n n' acc ->
-             match Node.Map.find n user_reads with
-             | exception Not_found -> acc
-             | domain ->
-               let domain' =
-                 try
-                   S.union (Node.Map.find n' acc) domain
-                 with Not_found -> domain
-               in
-               Node.Map.add n' domain' acc)
-          subst user_reads
-      in
-      (* For each node what must be read for assumes *)
-      let reads =
-        Bag.fold_left (fun acc e ->
-            Node.Map.union
-              (fun _ -> S.union) acc
-              (P.reads e))
-          user_reads env.assumes in
-      (* compute sigmas and relocate them *)
-      let env, sigmas = domains env reads pre in
-      if Wp_parameters.has_dkey dkey then
-        Format.printf "@[3) %a@]@." pretty_env env;
-      if Wp_parameters.has_dkey dumpkey then
-        dump_env ~name env;
-      let f, preds =
-        match mode with
-        | `Tree ->
-          (* Add a unique post node *)
-          let final_node = node () in
-          let env =
-            Node.Set.fold (fun p cfg ->
-                let s = {pre=S.create();post=S.create()} in
-                let e =  s,Lang.F.p_true in
-                let goto = memory_effect p e final_node in
-                concat goto cfg
-              ) posts env
-          in
-          To_tree.to_sequence_tree pre posts env
-        | (`Bool_Backward | `Bool_Forward) as mode ->
-          to_sequence_bool ~mode pre posts env in
-      let predssigmas =
-        Node.Map.merge
-          (fun _ p s -> Some (Option.value ~default:F.p_false p, Option.value ~default:(S.create ()) s))
-          preds sigmas in
-      (* readd simplified nodes *)
-      let predssigmas =
-        Node.Map.fold (fun n n' acc -> Node.Map.add n (Node.Map.find n' predssigmas) acc )
-          subst predssigmas
-      in
-      let preds =
-        Node.Map.map(fun _ (x,_) -> x) predssigmas
-      in
-      let sigmas =
-        Node.Map.map(fun _ (_,x) -> x) predssigmas
-      in
-      preds,sigmas,f
-
-end
+      Format.printf "@[3) %a@]@." pretty_env env;
+    if Wp_parameters.has_dkey dumpkey then
+      dump_env ~name env;
+    let f, preds =
+      match mode with
+      | `Tree ->
+        (* Add a unique post node *)
+        let final_node = node () in
+        let env =
+          Node.Set.fold (fun p cfg ->
+              let s = {pre=Sigma.create();post=Sigma.create()} in
+              let e =  s,Lang.F.p_true in
+              let goto = memory_effect p e final_node in
+              concat goto cfg
+            ) posts env
+        in
+        To_tree.to_sequence_tree pre posts env
+      | (`Bool_Backward | `Bool_Forward) as mode ->
+        to_sequence_bool ~mode pre posts env in
+    let predssigmas =
+      Node.Map.merge
+        (fun _ p s -> Some (
+             Option.value ~default:F.p_false p,
+             Option.value ~default:(Sigma.create ()) s))
+        preds sigmas in
+    (* readd simplified nodes *)
+    let predssigmas =
+      Node.Map.fold
+        (fun n n' acc -> Node.Map.add n (Node.Map.find n' predssigmas) acc )
+        subst predssigmas
+    in
+    let preds =
+      Node.Map.map(fun _ (x,_) -> x) predssigmas
+    in
+    let sigmas =
+      Node.Map.map(fun _ (_,x) -> x) predssigmas
+    in
+    preds,sigmas,f
diff --git a/src/plugins/wp/CfgCompiler.mli b/src/plugins/wp/CfgCompiler.mli
index bd0495a73c420b55574c13a986089af7ec2ce727..2d8ca02b5178a826024249c9aec10ec426e35bf7 100644
--- a/src/plugins/wp/CfgCompiler.mli
+++ b/src/plugins/wp/CfgCompiler.mli
@@ -20,7 +20,8 @@
 (*                                                                        *)
 (**************************************************************************)
 
-open Sigs
+open Memory
+open Sigma
 open Cil_types
 open Lang
 
@@ -51,233 +52,228 @@ type mode = [
   | `Bool_Forward
 ]
 
-module type Cfg =
+(** Program point along a trace. *)
+module Node : sig
+  type t
+  module Map : Qed.Idxmap.S with type key = t
+  module Set : Qed.Idxset.S with type elt = t
+  module Hashtbl : Hashtbl.S with type key = t
+  val pp: Format.formatter -> t -> unit
+  val create: unit -> t
+  val equal: t -> t -> bool
+end
+
+type node = Node.t
+
+(** fresh node *)
+val node : unit -> node
+
+(** {2 Relocatable Formulae}
+    Can be created once with fresh environment, and used several
+    times on different memory states. *)
+
+(** Relocatable condition *)
+module C :
+sig
+  type t
+
+  val equal : t -> t -> bool
+
+  (** Bundle an equation with the sigma sequence that created it. *)
+  val create : sigma -> F.pred -> t
+  val get : t -> F.pred
+  val reads : t -> domain
+  val relocate : sigma -> t -> t
+end
+
+(** Relocatable predicate *)
+module P :
 sig
+  type t
+  val pretty : Format.formatter -> t -> unit
 
-  (** The memory model used. *)
-  module S : Sigma
-
-  (** Program point along a trace. *)
-  module Node : sig
-    type t
-    module Map : Qed.Idxmap.S with type key = t
-    module Set : Qed.Idxset.S with type elt = t
-    module Hashtbl : Hashtbl.S with type key = t
-    val pp: Format.formatter -> t -> unit
-    val create: unit -> t
-    val equal: t -> t -> bool
-  end
-
-  type node = Node.t
-
-  (** fresh node *)
-  val node : unit -> node
-
-  (** {2 Relocatable Formulae}
-      Can be created once with fresh environment, and used several
-      times on different memory states. *)
-
-  (** Relocatable condition *)
-  module C :
-  sig
-    type t
-
-    val equal : t -> t -> bool
-
-    (** Bundle an equation with the sigma sequence that created it. *)
-    val create : S.t -> F.pred -> t
-    val get : t -> F.pred
-    val reads : t -> S.domain
-    val relocate : S.t -> t -> t
-  end
-
-  (** Relocatable predicate *)
-  module P :
-  sig
-    type t
-    val pretty : Format.formatter -> t -> unit
-
-    (** Bundle an equation with the sigma sequence that created it.
-        [| create m p |] = [| p |]
-    *)
-    val create : S.t Node.Map.t -> F.pred -> t
-    val get: t -> F.pred
-    val reads : t -> S.domain Node.Map.t
-    val nodes : t -> Node.Set.t
-    val relocate : S.t Node.Map.t -> t -> t
-    (** [| relocate m' (create m p) |] = [| p{ } |] *)
-
-    val to_condition: t -> (C.t * Node.t option) option
-  end
-
-  (** Relocatable term *)
-  module T :
-  sig
-    type t
-    val pretty : Format.formatter -> t -> unit
-
-    (** Bundle a term with the sigma sequence that created it. *)
-    val create : S.t Node.Map.t -> F.term -> t
-    val get: t -> F.term
-    val reads : t -> S.domain Node.Map.t
-    val relocate : S.t Node.Map.t -> t -> t
-    val init  : Node.Set.t -> (S.t Node.Map.t -> F.term) -> t
-    val init' : Node.t -> (S.t -> F.term) -> t
-  end
-
-
-  (** Relocatable effect (a predicate that depend on two states). *)
-  module E : sig
-    type t
-    val pretty: Format.formatter -> t -> unit
-
-    (** Bundle an equation with the sigma sequence that created it *)
-    val create : S.t sequence -> F.pred -> t
-    val get : t -> F.pred
-    val reads : t -> S.domain
-    val writes : t -> S.domain
-    (** as defined by S.writes *)
-
-    val relocate : S.t sequence -> t -> t
-  end
-
-  type cfg (** Structured collection of traces. *)
-
-  val dump_env: name:string -> cfg -> unit
-  val output_dot: out_channel -> ?checks:P.t Bag.t -> cfg -> unit
-
-  val nop : cfg
-  (** Structurally, [nop] is an empty execution trace.
-      Hence, [nop] actually denotes all possible execution traces.
-      This is the neutral element of [concat].
-
-      Formally: all interpretations I verify nop: [| nop |]_I
+  (** Bundle an equation with the sigma sequence that created it.
+      [| create m p |] = [| p |]
   *)
+  val create : sigma Node.Map.t -> F.pred -> t
+  val get: t -> F.pred
+  val reads : t -> domain Node.Map.t
+  val nodes : t -> Node.Set.t
+  val relocate : sigma Node.Map.t -> t -> t
+  (** [| relocate m' (create m p) |] = [| p{ } |] *)
+
+  val to_condition: t -> (C.t * Node.t option) option
+end
 
-  val add_tmpnode: node -> cfg
-  (** Set a node as temporary. Information about its path predicate or
-      sigma can be discarded during compilation *)
+(** Relocatable term *)
+module T :
+sig
+  type t
+  val pretty : Format.formatter -> t -> unit
+
+  (** Bundle a term with the sigma sequence that created it. *)
+  val create : sigma Node.Map.t -> F.term -> t
+  val get: t -> F.term
+  val reads : t -> domain Node.Map.t
+  val relocate : sigma Node.Map.t -> t -> t
+  val init  : Node.Set.t -> (sigma Node.Map.t -> F.term) -> t
+  val init' : Node.t -> (sigma -> F.term) -> t
+end
 
-  val concat : cfg -> cfg -> cfg
-  (** The concatenation is the intersection of all
-      possible collection of traces from each cfg.
 
-      [concat] is associative, commutative,
-      has [nop] as neutral element.
+(** Relocatable effect (a predicate that depend on two states). *)
+module E : sig
+  type t
+  val pretty: Format.formatter -> t -> unit
 
-      Formally: [| concat g1 g2 |]_I iff [| g1 |]_I and [| g2 |]_I
-  *)
+  (** Bundle an equation with the sigma sequence that created it *)
+  val create : sigma sequence -> F.pred -> t
+  val get : t -> F.pred
+  val reads : t -> domain
+  val writes : t -> domain
+  (** as defined by S.writes *)
 
-  val meta : ?stmt:stmt -> ?descr:string -> node -> cfg
-  (** Attach meta informations to a node.
-      Formally, it is equivalent to [nop]. *)
+  val relocate : sigma sequence -> t -> t
+end
 
-  val goto : node -> node -> cfg
-  (** Represents all execution traces [T] such that, if [T] contains node [a],
-      [T] also contains node [b] and memory states at [a] and [b] are equal.
+type cfg (** Structured collection of traces. *)
 
-      Formally: [| goto a b |]_I iff (I(a) iff I(b))
-  *)
+val dump_env: name:string -> cfg -> unit
+val output_dot: out_channel -> ?checks:P.t Bag.t -> cfg -> unit
 
-  val branch : node -> C.t -> node -> node -> cfg
-  (** Structurally corresponds to an if-then-else control-flow.
-      The predicate [P] shall reads only memory state at label [Here].
+val nop : cfg
+(** Structurally, [nop] is an empty execution trace.
+    Hence, [nop] actually denotes all possible execution traces.
+    This is the neutral element of [concat].
 
-      Formally: [| branch n P a b |]_I iff (   (I(n) iff (I(a) \/ I(b)))
-                                            /\ (I(n) implies (if P(I(n)) then I(a) else I(b)))  )
-  *)
+    Formally: all interpretations I verify nop: [| nop |]_I
+*)
 
-  val guard : node -> C.t -> node -> cfg
-  (** Structurally corresponds to an assume control-flow.
-      The predicate [P] shall reads only memory state at label [Here].
+val add_tmpnode: node -> cfg
+(** Set a node as temporary. Information about its path predicate or
+    sigma can be discarded during compilation *)
 
-      Formally: [| guard n P a |]_I iff (   (I(n) iff I(a))
-                                            /\ (I(n) implies [| P |]_I  ) )
-  *)
+val concat : cfg -> cfg -> cfg
+(** The concatenation is the intersection of all
+    possible collection of traces from each cfg.
 
-  val guard' : node -> C.t -> node -> cfg
-  (** Same than guard but the condition is negated *)
+    [concat] is associative, commutative,
+    has [nop] as neutral element.
 
-  val either : node -> node list -> cfg
-  (** Structurally corresponds to an arbitrary choice among the different
-      possible executions.
+    Formally: [| concat g1 g2 |]_I iff [| g1 |]_I and [| g2 |]_I
+*)
 
-      [either] is associative and commutative. [either a []] is
-      very special, since it denotes a cfg with {i no} trace. Technically,
-      it is equivalent to attaching an [assert \false] annotation to node [a].
+val meta : ?stmt:stmt -> ?descr:string -> node -> cfg
+(** Attach meta informations to a node.
+    Formally, it is equivalent to [nop]. *)
 
-      Formally: [| either n [a_1;...;a_n] } |]_I iff ( I(n) iff (I(a_1) \/ ... I(a_n)))
-  *)
+val goto : node -> node -> cfg
+(** Represents all execution traces [T] such that, if [T] contains node [a],
+    [T] also contains node [b] and memory states at [a] and [b] are equal.
 
+    Formally: [| goto a b |]_I iff (I(a) iff I(b))
+*)
 
-  val implies : node -> (C.t * node) list -> cfg
-  (**
-     implies is the dual of either. Instead of being a non-deterministic choice,
-     it takes the choices that verify its predicate.
+val branch : node -> C.t -> node -> node -> cfg
+(** Structurally corresponds to an if-then-else control-flow.
+    The predicate [P] shall reads only memory state at label [Here].
 
-      Formally: [| either n [P_1,a_1;...;P_n,a_n] } |]_I iff ( I(n) iff (I(a_1) \/ ... I(a_n))
-                                                              /\  I(n) implies [| P_k |]_I implies I(a_k)
-  *)
+    Formally: [| branch n P a b |]_I iff (   (I(n) iff (I(a) \/ I(b)))
+                                          /\ (I(n) implies (if P(I(n)) then I(a) else I(b)))  )
+*)
 
+val guard : node -> C.t -> node -> cfg
+(** Structurally corresponds to an assume control-flow.
+    The predicate [P] shall reads only memory state at label [Here].
 
-  val memory_effect : node -> E.t -> node -> cfg
-  (** Represents all execution trace [T] such that, if [T] contains node [a],
-      then [T] also contains [b] with the given effect on corresponding
-      memory states.
+    Formally: [| guard n P a |]_I iff (   (I(n) iff I(a))
+                                          /\ (I(n) implies [| P |]_I  ) )
+*)
 
-      Formally: [| memory_effect a e b |]_I iff (( I(a) iff I(b) ) /\ [| e |]_I )
-  *)
+val guard' : node -> C.t -> node -> cfg
+(** Same than guard but the condition is negated *)
 
-  val assume : P.t -> cfg
-  (** Represents execution traces [T] such that, if [T] contains
-      every node points in the label-map, then the condition holds over the
-      corresponding memory states. If the node-map is empty,
-      the condition must hold over all possible execution path.
+val either : node -> node list -> cfg
+(** Structurally corresponds to an arbitrary choice among the different
+    possible executions.
 
-      Formally: [| assume P |]_I iff [| P |]_I
-  *)
+    [either] is associative and commutative. [either a []] is
+    very special, since it denotes a cfg with {i no} trace. Technically,
+    it is equivalent to attaching an [assert \false] annotation to node [a].
 
-  val havoc : node -> memory_effects:node sequence -> node -> cfg
-  (** Inserts an assigns effect between nodes [a] and [b], correspondings
-      to all the written memory chunks accessible in execution paths delimited
-      by the [effects] sequence of nodes.
+    Formally: [| either n [a_1;...;a_n] } |]_I iff ( I(n) iff (I(a_1) \/ ... I(a_n)))
+*)
 
-      Formally: [| havoc a s b |]_I is verified if there is no path between s.pre and s.path,
-      otherwise if (I(a) iff I(b) and if I(a) is defined then I(a) and I(b) are equal
-      for all the chunks that are not in the written domain of an effect that can be found
-      between [s.pre] to [s.post].
 
-      Note: the effects are collected in the {i final} control-flow,
-      when {!compile} is invoked. The portion of the sub-graph in the sequence
-      shall be concatenated to the [cfg] before compiling-it, otherwize it would be
-      considered empty and [havoc] would be a nop (no connection between a and b).
-  *)
+val implies : node -> (C.t * node) list -> cfg
+(**
+   implies is the dual of either. Instead of being a non-deterministic choice,
+   it takes the choices that verify its predicate.
 
-  (** {2 Path-Predicates}
+    Formally: [| either n [P_1,a_1;...;P_n,a_n] } |]_I iff ( I(n) iff (I(a_1) \/ ... I(a_n))
+                                                            /\  I(n) implies [| P_k |]_I implies I(a_k)
+*)
 
-      The compilation of cfg control-flow into path predicate
-      is performed by allocating fresh environments with optimized variable
-      allocation. Only the relevant path between the nodes
-      is extracted. Other paths in the cfg are pruned out.
-  *)
 
-  (** Extract the nodes that are between the start node and the final
-      nodes and returns how to observe a collection of states indexed
-      by nodes. The returned maps gives, for each reachable node, a
-      predicate representing paths that reach the node and the memory
-      state at this node.
+val memory_effect : node -> E.t -> node -> cfg
+(** Represents all execution trace [T] such that, if [T] contains node [a],
+    then [T] also contains [b] with the given effect on corresponding
+    memory states.
 
-      Nodes absent from the map are unreachable. Whenever possible,
-      predicate [F.ptrue] is returned for inconditionally accessible
-      nodes.
+    Formally: [| memory_effect a e b |]_I iff (( I(a) iff I(b) ) /\ [| e |]_I )
+*)
 
-      ~name: identifier used for debugging
+val assume : P.t -> cfg
+(** Represents execution traces [T] such that, if [T] contains
+    every node points in the label-map, then the condition holds over the
+    corresponding memory states. If the node-map is empty,
+    the condition must hold over all possible execution path.
 
-  *)
+    Formally: [| assume P |]_I iff [| P |]_I
+*)
 
-  val compile : ?name:string -> ?mode:mode -> node -> Node.Set.t -> S.domain Node.Map.t ->
-    cfg -> F.pred Node.Map.t * S.t Node.Map.t * Conditions.sequence
+val havoc : node -> memory_effects:node sequence -> node -> cfg
+(** Inserts an assigns effect between nodes [a] and [b], correspondings
+    to all the written memory chunks accessible in execution paths delimited
+    by the [effects] sequence of nodes.
 
-end
+    Formally: [| havoc a s b |]_I is verified if there is no path between s.pre and s.path,
+    otherwise if (I(a) iff I(b) and if I(a) is defined then I(a) and I(b) are equal
+    for all the chunks that are not in the written domain of an effect that can be found
+    between [s.pre] to [s.post].
+
+    Note: the effects are collected in the {i final} control-flow,
+    when {!compile} is invoked. The portion of the sub-graph in the sequence
+    shall be concatenated to the [cfg] before compiling-it, otherwize it would be
+    considered empty and [havoc] would be a nop (no connection between a and b).
+*)
+
+(** {2 Path-Predicates}
+
+    The compilation of cfg control-flow into path predicate
+    is performed by allocating fresh environments with optimized variable
+    allocation. Only the relevant path between the nodes
+    is extracted. Other paths in the cfg are pruned out.
+*)
+
+(** Extract the nodes that are between the start node and the final
+    nodes and returns how to observe a collection of states indexed
+    by nodes. The returned maps gives, for each reachable node, a
+    predicate representing paths that reach the node and the memory
+    state at this node.
+
+    Nodes absent from the map are unreachable. Whenever possible,
+    predicate [F.ptrue] is returned for inconditionally accessible
+    nodes.
+
+    ~name: identifier used for debugging
+
+*)
 
-module Cfg(S:Sigma) : Cfg with module S = S
+val compile :
+  ?name:string -> ?mode:mode ->
+  node ->
+  Node.Set.t ->
+  domain Node.Map.t ->
+  cfg ->
+  F.pred Node.Map.t * sigma Node.Map.t * Conditions.sequence
diff --git a/src/plugins/wp/CodeSemantics.ml b/src/plugins/wp/CodeSemantics.ml
index 705c50ae21ab1ee8f3c104f75af0776929d35951..7425baceb102adc967cb876f4412bd4e37d378cb 100644
--- a/src/plugins/wp/CodeSemantics.ml
+++ b/src/plugins/wp/CodeSemantics.ml
@@ -28,9 +28,9 @@ open Cil_datatype
 open Cil_types
 open Ctypes
 open Qed
-open Sigs
+open Memory
+open Sigma
 open Lang
-open Lang.F
 
 module WpLog = Wp_parameters
 let constfold_ctyp = function
@@ -52,15 +52,14 @@ let constfold_coffset = function
     end
   | off -> off
 
-module Make(M : Sigs.Model) =
+module Make(M : Memory.Model) =
 struct
 
   module M = M
 
   type loc = M.loc
-  type value = M.loc Sigs.value
-  type sigma = M.Sigma.t
-  type result = loc Sigs.result
+  type value = M.loc Memory.value
+  type result = loc Memory.result
 
   let pp_value fmt = function
     | Val e -> Format.fprintf fmt "Val:%a" F.pp_term e
@@ -79,11 +78,11 @@ struct
   (* -------------------------------------------------------------------------- *)
 
   let is_zero_int = function
-    | Val e -> p_equal e e_zero
+    | Val e -> F.p_equal e F.e_zero
     | Loc l -> M.is_null l
 
   let is_zero_float ft = function
-    | Val e -> p_equal e @@ Cfloat.float_of_real ft e_zero_real
+    | Val e -> F.p_equal e @@ Cfloat.float_of_real ft F.e_zero_real
     | Loc l -> M.is_null l
 
   let is_zero_ptr v = M.is_null (cloc v)
@@ -94,9 +93,9 @@ struct
     | C_float ft -> is_zero_float ft (M.load sigma obj l)
     | C_pointer _ -> is_zero_ptr (M.load sigma obj l)
     | C_comp { cfields = None } ->
-      p_true (* cannot say anything interesting here *)
+      F.p_true (* cannot say anything interesting here *)
     | C_comp { cfields = Some fields } ->
-      p_all
+      F.p_all
         (fun f -> is_zero sigma (Ctypes.object_of f.ftype) (M.field l f))
         fields
     | C_array a ->
@@ -104,18 +103,18 @@ struct
                  For instance, a[N][M] becomes a[N*M] in MemTyped,
                  but not in MemVar *)
       let x = Lang.freshvar ~basename:"k" Logic.Int in
-      let k = e_var x in
+      let k = F.e_var x in
       let obj = Ctypes.object_of a.arr_element in
       let range = match a.arr_flat with
         | None -> []
-        | Some f -> [ p_leq e_zero k ; p_lt k (e_int f.arr_size) ] in
+        | Some f -> [ F.p_leq F.e_zero k ; F.p_lt k (F.e_int f.arr_size) ] in
       let init = is_zero sigma obj (M.shift l obj k) in
-      p_forall [x] (p_hyps range init)
+      F.p_forall [x] (F.p_hyps range init)
 
   let is_exp_range sigma l obj a b v =
     let x = Lang.freshvar ~basename:"k" Logic.Int in
-    let k = e_var x in
-    let range = [ p_leq a k ; p_leq k b ] in
+    let k = F.e_var x in
+    let range = [ F.p_leq a k ; F.p_leq k b ] in
     let init =
       match v with
       | None -> is_zero sigma obj (M.shift l obj k)
@@ -124,16 +123,16 @@ struct
         if Ctypes.is_pointer obj then
           M.loc_eq (cloc elt) (cloc v)
         else
-          p_equal (cval elt) (cval v)
+          F.p_equal (cval elt) (cval v)
     in
-    p_forall [x] (p_hyps range init)
+    F.p_forall [x] (F.p_hyps range init)
 
   (* -------------------------------------------------------------------------- *)
   (* --- Recursion                                                          --- *)
   (* -------------------------------------------------------------------------- *)
 
   let s_exp : (sigma -> exp -> value) ref = ref (fun _ _ -> assert false)
-  let s_cond : (sigma -> exp -> pred) ref = ref (fun _ _ -> assert false)
+  let s_cond : (sigma -> exp -> F.pred) ref = ref (fun _ _ -> assert false)
 
   let val_of_exp env e = cval (!s_exp env e)
   let loc_of_exp env e = cloc (!s_exp env e)
@@ -168,8 +167,8 @@ struct
       | C_int i , Neg -> Cint.iopp i (val_of_exp env e)
       | C_int i , BNot -> Cint.bnot i (val_of_exp env e)
       | C_float f , Neg -> Cfloat.fopp f (val_of_exp env e)
-      | C_int _ , LNot -> Cvalues.bool_eq (val_of_exp env e) e_zero
-      | C_float _ , LNot -> Cvalues.bool_eq (val_of_exp env e) e_zero_real
+      | C_int _ , LNot -> Cvalues.bool_eq (val_of_exp env e) F.e_zero
+      | C_float _ , LNot -> Cvalues.bool_eq (val_of_exp env e) F.e_zero_real
       | C_pointer _ , LNot -> Cvalues.is_true (M.is_null (loc_of_exp env e))
       | _ ->
         Warning.error "Undefined unary operator (%a)" Printer.pp_typ typ
@@ -199,14 +198,14 @@ struct
       | TFloat(f,_) ->
         let p = fop (Ctypes.c_float f)
             (val_of_exp env e1) (val_of_exp env e2) in
-        e_if (F.e_prop p) e_one e_zero
+        F.e_if (F.e_prop p) F.e_one F.e_zero
       | _ ->
         iop (val_of_exp env e1) (val_of_exp env e2)
 
   let bool_of_exp env e =
     match Ctypes.object_of (Cil.typeOf e) with
-    | C_int _ -> Cvalues.bool_neq (val_of_exp env e) e_zero
-    | C_float _ -> Cvalues.bool_neq (val_of_exp env e) e_zero_real
+    | C_int _ -> Cvalues.bool_neq (val_of_exp env e) F.e_zero
+    | C_float _ -> Cvalues.bool_neq (val_of_exp env e) F.e_zero_real
     | C_pointer _ -> Cvalues.is_false (M.is_null (loc_of_exp env e))
     | _ -> assert false
 
@@ -236,7 +235,7 @@ struct
     | MinusPI ->
       let te = Cil.typeOf_pointed (Cil.typeOf e1) in
       let obj = Ctypes.object_of te in
-      Loc(M.shift (loc_of_exp env e1) obj (e_opp (val_of_exp env e2)))
+      Loc(M.shift (loc_of_exp env e1) obj (F.e_opp (val_of_exp env e2)))
     | MinusPP ->
       let te = Cil.typeOf_pointed (Cil.typeOf e1) in
       let obj = Ctypes.object_of te in
@@ -292,7 +291,7 @@ struct
   let exp_undefined e =
     let ty = Cil.typeOf e in
     let x = Lang.freshvar ~basename:"w" (Lang.tau_of_ctype ty) in
-    Val (e_var x)
+    Val (F.e_var x)
 
   (* -------------------------------------------------------------------------- *)
   (* --- Exp-Node                                                           --- *)
@@ -355,20 +354,20 @@ struct
   let eq_t is_ptr t v1 v2 =
     match v1 , v2 with
     | Loc p , Loc q -> M.loc_eq p q
-    | Val a , Val b -> p_equal a b
+    | Val a , Val b -> F.p_equal a b
     | _ ->
       if is_ptr t
       then M.loc_eq (cloc v1) (cloc v2)
-      else p_equal (cval v1) (cval v2)
+      else F.p_equal (cval v1) (cval v2)
 
   let neq_t is_ptr t v1 v2 =
     match v1 , v2 with
     | Loc p , Loc q -> M.loc_neq p q
-    | Val a , Val b -> p_neq a b
+    | Val a , Val b -> F.p_neq a b
     | _ ->
       if is_ptr t
       then M.loc_neq (cloc v1) (cloc v2)
-      else p_neq (cval v1) (cval v2)
+      else F.p_neq (cval v1) (cval v2)
 
   let equal_typ t v1 v2 = eq_t Cil.isPointerType t v1 v2
   let equal_obj obj v1 v2 = eq_t Ctypes.is_pointer obj v1 v2
@@ -389,22 +388,22 @@ struct
   let cond_node env e =
     match e.enode with
 
-    | UnOp(  LNot, e,_)     -> p_not (!s_cond env e)
-    | BinOp( LAnd, e1,e2,_) -> p_and (!s_cond env e1) (!s_cond env e2)
-    | BinOp( LOr,  e1,e2,_) -> p_or (!s_cond env e1) (!s_cond env e2)
-    | BinOp( Eq,   e1,e2,_) -> compare env p_equal M.loc_eq Cfloat.feq e1 e2
-    | BinOp( Ne,   e1,e2,_) -> compare env p_neq M.loc_neq Cfloat.fneq e1 e2
-    | BinOp( Lt,   e1,e2,_) -> compare env p_lt  M.loc_lt  Cfloat.flt e1 e2
-    | BinOp( Gt,   e1,e2,_) -> compare env p_lt  M.loc_lt  Cfloat.flt e2 e1
-    | BinOp( Le,   e1,e2,_) -> compare env p_leq M.loc_leq Cfloat.fle e1 e2
-    | BinOp( Ge,   e1,e2,_) -> compare env p_leq M.loc_leq Cfloat.fle e2 e1
+    | UnOp(  LNot, e,_)     -> F.p_not (!s_cond env e)
+    | BinOp( LAnd, e1,e2,_) -> F.p_and (!s_cond env e1) (!s_cond env e2)
+    | BinOp( LOr,  e1,e2,_) -> F.p_or (!s_cond env e1) (!s_cond env e2)
+    | BinOp( Eq,   e1,e2,_) -> compare env F.p_equal M.loc_eq Cfloat.feq e1 e2
+    | BinOp( Ne,   e1,e2,_) -> compare env F.p_neq M.loc_neq Cfloat.fneq e1 e2
+    | BinOp( Lt,   e1,e2,_) -> compare env F.p_lt  M.loc_lt  Cfloat.flt e1 e2
+    | BinOp( Gt,   e1,e2,_) -> compare env F.p_lt  M.loc_lt  Cfloat.flt e2 e1
+    | BinOp( Le,   e1,e2,_) -> compare env F.p_leq M.loc_leq Cfloat.fle e1 e2
+    | BinOp( Ge,   e1,e2,_) -> compare env F.p_leq M.loc_leq Cfloat.fle e2 e1
 
     | _ ->
       begin
         match Ctypes.object_of (Cil.typeOf e) with
-        | C_int _ -> p_neq (val_of_exp env e) e_zero
-        | C_float _ -> p_neq (val_of_exp env e) e_zero_real
-        | C_pointer _ -> p_not (M.is_null (loc_of_exp env e))
+        | C_int _ -> F.p_neq (val_of_exp env e) F.e_zero
+        | C_float _ -> F.p_neq (val_of_exp env e) F.e_zero_real
+        | C_pointer _ -> F.p_not (M.is_null (loc_of_exp env e))
         | obj -> Warning.error "Condition from (%a)" Ctypes.pretty obj
       end
 
@@ -446,14 +445,14 @@ struct
            let value_hyp = match init with
              | Some e ->
                let v = M.load sigma obj l in
-               p_equal (val_of_exp sigma e) (cval v)
+               F.p_equal (val_of_exp sigma e) (cval v)
              | None -> is_zero sigma obj l
            in
            let init_hyp = match init with
              | Some { enode = Lval lv_init }
                when Cil.(isStructOrUnionType @@ typeOfLval lv_init) ->
                let l_initializer = lval sigma lv_init in
-               p_equal
+               F.p_equal
                  (M.load_init sigma obj l)
                  (M.load_init sigma obj l_initializer)
              | _ ->
@@ -472,7 +471,7 @@ struct
         (fun () ->
            let l = lval sigma lv in
            let e = Option.map (exp sigma) value in
-           let low = e_bigint low and up = e_bigint up in
+           let low = F.e_bigint low and up = F.e_bigint up in
            (is_exp_range sigma l obj low up e),
            (M.initialized sigma (Rrange(l, obj, Some low, Some up)))
         ) () in
diff --git a/src/plugins/wp/CodeSemantics.mli b/src/plugins/wp/CodeSemantics.mli
index a3c32bb77d402f0deed3703e3219f5cf7f85cb52..05db16bbdb391f4c98a136ba11c643708e074215 100644
--- a/src/plugins/wp/CodeSemantics.mli
+++ b/src/plugins/wp/CodeSemantics.mli
@@ -25,4 +25,4 @@
 (* -------------------------------------------------------------------------- *)
 
 
-module Make(M : Sigs.Model) : Sigs.CodeSemantics with module M = M
+module Make(M : Memory.Model) : Memory.CodeSemantics with module M = M
diff --git a/src/plugins/wp/Conditions.ml b/src/plugins/wp/Conditions.ml
index dd63ee019bf5177ec2cf68666fa6a37ac603ecac..610aea99227c884a0c2f2f7afe6ed50b4565d39e 100644
--- a/src/plugins/wp/Conditions.ml
+++ b/src/plugins/wp/Conditions.ml
@@ -783,7 +783,7 @@ let ground (hs,g) =
 (* --- Letify                                                             --- *)
 (* -------------------------------------------------------------------------- *)
 
-module Sigma = Letify.Sigma
+module Subst = Letify.Subst
 module Defs = Letify.Defs
 
 let used_of_dseq ds =
@@ -801,12 +801,12 @@ let locals sigma ~target ~required ?(step=Vars.empty) k dseq =
        if i > k then t := Vars.union !t step.vars ;
        if i <> k then e := Vars.union !e step.vars ;
     ) dseq ;
-  Vars.diff !t (Sigma.domain sigma) , !e
+  Vars.diff !t (Subst.domain sigma) , !e
 
 let dseq_of_step sigma step =
   let defs =
     match step.condition with
-    | Init p | Have p | When p | Core p -> Defs.extract (Sigma.p_apply sigma p)
+    | Init p | Have p | When p | Core p -> Defs.extract (Subst.p_apply sigma p)
     | Type _ | Branch _ | Either _ | State _ | Probe _ -> Defs.empty
   in defs , step
 
@@ -817,13 +817,13 @@ let letify_assume sref (_,step) =
     | Type _ | Branch _ | Either _ | State _ | Probe _ -> ()
     | Init p | Have p | When p | Core p ->
       if Wp_parameters.Simpl.get () then
-        sref := Sigma.assume current p
+        sref := Subst.assume current p
   end ; current
 
 let rec letify_type sigma used p = match F.p_expr p with
   | And ps -> p_all (letify_type sigma used) ps
   | _ ->
-    let p = Sigma.p_apply sigma p in
+    let p = Subst.p_apply sigma p in
     let vs = F.varsp p in
     if Vars.intersect used vs || Vars.is_empty vs then p else F.p_true
 
@@ -835,11 +835,11 @@ let rec letify_seq sigma0 ~target ~export (seq : step list) =
   let sigma2 = !sref in (* with assumptions *)
   let outside = Vars.union export target in
   let inside = used_of_dseq dseq in
-  let used = Vars.diff (Vars.union outside inside) (Sigma.domain sigma2) in
-  let required = Vars.union outside (Sigma.codomain sigma2) in
+  let used = Vars.diff (Vars.union outside inside) (Subst.domain sigma2) in
+  let required = Vars.union outside (Subst.codomain sigma2) in
   let sequence =
     Array.mapi (letify_step dseq dsigma ~used ~required ~target) dseq in
-  let modified = ref (not (Sigma.equal sigma0 sigma1)) in
+  let modified = ref (not (Subst.equal sigma0 sigma1)) in
 (*
   let sequence =
     if Wp_parameters.Ground.get () then fst (ground_hrp sequence)
@@ -851,32 +851,32 @@ let rec letify_seq sigma0 ~target ~export (seq : step list) =
 and letify_step dseq dsigma ~required ~target ~used i (d,s) =
   let sigma = dsigma.(i) in
   let cond = match s.condition with
-    | Probe(p,t) -> Probe (p,Sigma.e_apply sigma t)
-    | State s -> State (Mstate.apply (Sigma.e_apply sigma) s)
+    | Probe(p,t) -> Probe (p,Subst.e_apply sigma t)
+    | State s -> State (Mstate.apply (Subst.e_apply sigma) s)
     | Init p ->
-      let p = Sigma.p_apply sigma p in
+      let p = Subst.p_apply sigma p in
       let ps = Letify.add_definitions sigma d required [p] in
       Init (p_conj ps)
     | Have p ->
-      let p = Sigma.p_apply sigma p in
+      let p = Subst.p_apply sigma p in
       let ps = Letify.add_definitions sigma d required [p] in
       Have (p_conj ps)
     | Core p ->
-      let p = Sigma.p_apply sigma p in
+      let p = Subst.p_apply sigma p in
       let ps = Letify.add_definitions sigma d required [p] in
       Core (p_conj ps)
     | When p ->
-      let p = Sigma.p_apply sigma p in
+      let p = Subst.p_apply sigma p in
       let ps = Letify.add_definitions sigma d required [p] in
       When (p_conj ps)
     | Type p ->
       Type (letify_type sigma used p)
     | Branch(p,a,b) ->
-      let p = Sigma.p_apply sigma p in
+      let p = Subst.p_apply sigma p in
       let step = F.varsp p in
       let (target,export) = locals sigma ~target ~required ~step i dseq in
-      let sa = Sigma.assume sigma p in
-      let sb = Sigma.assume sigma (p_not p) in
+      let sa = Subst.assume sigma p in
+      let sb = Subst.assume sigma (p_not p) in
       let a = letify_case sa ~target ~export a in
       let b = letify_case sb ~target ~export b in
       Branch(p,a,b)
@@ -1173,7 +1173,7 @@ and compute limit solvers sigma s0 =
   let export = Vars.empty in
   let modified , sigma1 , sigma2 , hs =
     letify_seq sigma ~target ~export hs in
-  let p = Sigma.p_apply sigma2 p in
+  let p = Subst.p_apply sigma2 p in
   let s2 = ground (hs , p) in
   if is_trivial_hsp s2 then [],p_true
   else
@@ -1185,10 +1185,10 @@ and compute limit solvers sigma s0 =
     | Trivial -> [],p_true
     | NoSimplification -> s2
 
-let letify_hsp ?(solvers=[]) hsp = fixpoint 10 solvers Sigma.empty hsp
+let letify_hsp ?(solvers=[]) hsp = fixpoint 10 solvers Subst.empty hsp
 
 let rec simplify ?(solvers=[]) ?(intros=10) (seq,p0) =
-  let hs,p = fixpoint 10 solvers Sigma.empty (seq.seq_list,p0) in
+  let hs,p = fixpoint 10 solvers Subst.empty (seq.seq_list,p0) in
   let sequent = sequence hs , p in
   match introduction sequent with
   | Some introduced ->
@@ -1685,8 +1685,9 @@ struct
         if Lang.F.is_primitive t then ()
         else
           match Mstate.lookup s t with
-          | Mlval(_, KInit) | Mchunk(_, KInit) -> raise Found
-          | _ -> ()
+          | Minit _ -> raise Found
+          | Mchunk chunk when Sigma.is_init chunk -> raise Found
+          | exception Not_found | _ -> ()
       in
       if Tset.mem t !visited then ()
       else begin
diff --git a/src/plugins/wp/Cvalues.ml b/src/plugins/wp/Cvalues.ml
index cf0e615fb34bbe2bca602a4dba79726e9b3bfc61..f57fd630a5c5eb46781a6d21be88d3e5a54a6c35 100644
--- a/src/plugins/wp/Cvalues.ml
+++ b/src/plugins/wp/Cvalues.ml
@@ -29,7 +29,7 @@ open Ctypes
 open Qed
 open Lang
 open Lang.F
-open Sigs
+open Memory
 open Definitions
 
 (* -------------------------------------------------------------------------- *)
@@ -583,12 +583,12 @@ let negate = function
   | `Negative -> `Positive
   | `NoPolarity -> `NoPolarity
 
-module Logic(M : Sigs.Model) =
+module Logic(M : Memory.Model) =
 struct
 
-  type logic = M.loc Sigs.logic
-  type segment = c_object * M.loc Sigs.sloc
-  type region = M.loc Sigs.region
+  type logic = M.loc Memory.logic
+  type segment = c_object * M.loc Memory.sloc
+  type region = M.loc Memory.region
 
   (* -------------------------------------------------------------------------- *)
   (* --- Projections                                                        --- *)
diff --git a/src/plugins/wp/Cvalues.mli b/src/plugins/wp/Cvalues.mli
index c21b6053e152cdc1ed69e78b85260b6786cd0de5..a1817a39962dfe4fbc850de7f5f19b923332e507 100644
--- a/src/plugins/wp/Cvalues.mli
+++ b/src/plugins/wp/Cvalues.mli
@@ -26,10 +26,11 @@
 
 open Cil_types
 open Ctypes
-open Sigs
 open Lang.F
+open Memory
+open Sigma
 
-val equation : Sigs.equation -> pred
+val equation : Memory.equation -> pred
 
 (** {2 Pretty Printing} *)
 
@@ -113,10 +114,10 @@ val bytes_length_of_opaque_comp: compinfo -> term
 
 (** {2 Lifting Operations over Memory Values} *)
 
-val map_sloc : ('a -> 'b) -> 'a Sigs.sloc -> 'b Sigs.sloc
-val map_value : ('a -> 'b) -> 'a Sigs.value -> 'b Sigs.value
-val map_logic : ('a -> 'b) -> 'a Sigs.logic -> 'b Sigs.logic
-val plain : logic_type -> term -> 'a Sigs.logic
+val map_sloc : ('a -> 'b) -> 'a Memory.sloc -> 'b Memory.sloc
+val map_value : ('a -> 'b) -> 'a Memory.value -> 'b Memory.value
+val map_logic : ('a -> 'b) -> 'a Memory.logic -> 'b Memory.logic
+val plain : logic_type -> term -> 'a Memory.logic
 
 (** {2 ACSL Utilities} *)
 
@@ -126,13 +127,13 @@ val plain : logic_type -> term -> 'a Sigs.logic
 type polarity = [ `Positive | `Negative | `NoPolarity ]
 val negate : polarity -> polarity
 
-module Logic(M : Sigs.Model) :
+module Logic(M : Memory.Model) :
 sig
 
   open M
-  type logic = M.loc Sigs.logic
-  type segment = c_object * loc Sigs.sloc
-  type region = loc Sigs.region
+  type logic = M.loc Memory.logic
+  type segment = c_object * loc Memory.sloc
+  type region = loc Memory.region
 
   (** {3 Projections} *)
 
@@ -158,7 +159,7 @@ sig
 
   val field : logic -> fieldinfo -> logic
   val shift : logic -> c_object -> ?size:int -> logic -> logic
-  val load : Sigma.t -> c_object -> logic -> logic
+  val load : sigma -> c_object -> logic -> logic
 
   (** {3 Sets of loc-or-values} *)
 
@@ -170,8 +171,8 @@ sig
 
   val separated : region list -> pred
   val included : segment -> segment -> pred
-  val valid : Sigma.t -> acs -> segment -> pred
-  val invalid : Sigma.t -> segment -> pred
-  val initialized : Sigma.t -> segment -> pred
+  val valid : sigma -> acs -> segment -> pred
+  val invalid : sigma -> segment -> pred
+  val initialized : sigma -> segment -> pred
 
 end
diff --git a/src/plugins/wp/Factory.ml b/src/plugins/wp/Factory.ml
index bc4668dbe460cc70f87d2e4f1474324ba061fb79..3b077a816ce1d499dc88c0ac0303ddf9cd995a0f 100644
--- a/src/plugins/wp/Factory.ml
+++ b/src/plugins/wp/Factory.ml
@@ -204,7 +204,7 @@ module MemEva = MemVal.Make(MemVal.Eva)
 (* Each model must be instanciated statically because of registered memory
    models identifiers and Frama-C states *)
 
-module MakeVar(V : MemVar.VarUsage)(M : Sigs.Model)
+module MakeVar(V : MemVar.VarUsage)(M : Memory.Model)
   = MemVar.Make(MakeVarUsage(V))(M)
 
 module Model_Hoare_Raw = MakeVar(MemVar.Raw)(MemEmpty)
@@ -216,7 +216,7 @@ module Model_Bytes_Ref = MakeVar(Ref)(MemBytes)
 module Model_Caveat = MakeVar(Caveat)(MemTyped)
 module Model_Region = MakeVar(Static)(MemRegion)
 
-module MakeCompiler(M:Sigs.Model) =
+module MakeCompiler(M:Memory.Model) =
 struct
   module M = M
   module C = CodeSemantics.Make(M)
@@ -239,7 +239,7 @@ module Comp_Bytes_Ref = MakeCompiler(Model_Bytes_Ref)
 module Comp_Region = MakeCompiler(Model_Region)
 module Comp_MemEva = MakeCompiler(MemEva)
 
-let compiler mheap mvar : (module Sigs.Compiler) =
+let compiler mheap mvar : (module Memory.Compiler) =
   match mheap , mvar with
   | ZeroAlias , _     -> (module Comp_MemZeroAlias)
   | _    ,   Caveat   -> (module Comp_Caveat)
diff --git a/src/plugins/wp/Factory.mli b/src/plugins/wp/Factory.mli
index 50192a0e924f1896dd359dabec1e49a6adfbb42c..b138dc1cbebbd8b1f3a547340299b2a18892f047 100644
--- a/src/plugins/wp/Factory.mli
+++ b/src/plugins/wp/Factory.mli
@@ -38,7 +38,7 @@ type driver = LogicBuiltins.driver
 
 val ident : setup -> string
 val descr : setup -> string
-val compiler : mheap -> mvar -> (module Sigs.Compiler)
+val compiler : mheap -> mvar -> (module Memory.Compiler)
 val configure_driver : setup -> driver -> unit -> WpContext.rollback
 val instance : setup -> driver -> WpContext.model
 val default : setup (** ["Var,Typed,Nat,Real"] memory model. *)
diff --git a/src/plugins/wp/Letify.ml b/src/plugins/wp/Letify.ml
index 45aeafcebaaf6f50910e5930acd8be71a3bd547d..8a665bd0e9d6ddea81d13e71755ebb7eeb513afd 100644
--- a/src/plugins/wp/Letify.ml
+++ b/src/plugins/wp/Letify.ml
@@ -148,7 +148,7 @@ end
 (* --- Generalized Substitution                                           --- *)
 (* -------------------------------------------------------------------------- *)
 
-module Sigma :
+module Subst :
 sig
   type t
   val equal : t -> t -> bool
@@ -488,7 +488,7 @@ let elements xs = Vars.fold XS.add xs XS.empty
 let iter f xs = XS.iter f (elements xs)
 
 let rec extract defs sref cycle x =
-  if not (Vars.mem x cycle) && not (Sigma.mem x !sref) then
+  if not (Vars.mem x cycle) && not (Subst.mem x !sref) then
     try
       let cycle = Vars.add x cycle in
       let ds = Vmap.find x defs in (* if no defs, exit early *)
@@ -501,13 +501,13 @@ let rec extract defs sref cycle x =
              match F.repr e with
              | Fvar y ->
                begin
-                 try let d = Sigma.find y !sref in rs := d :: !rs
+                 try let d = Subst.find y !sref in rs := d :: !rs
                  with Not_found -> ys := y :: !ys
                end
              | _ -> es := e :: !es
         ) ds ;
       (* Now choose the represent of x and the dependencies *)
-      let select d = sref := Sigma.add x d !sref ; d , F.vars d in
+      let select d = sref := Subst.add x d !sref ; d , F.vars d in
       let ceq , depends =
         match List.sort F.compare !rs with
         | r :: _ -> select r
@@ -515,7 +515,7 @@ let rec extract defs sref cycle x =
           | e :: _ -> select e
           | [] -> e_var x , Vars.empty
       in
-      List.iter (fun y -> sref := Sigma.add y ceq !sref) !ys ;
+      List.iter (fun y -> sref := Subst.add y ceq !sref) !ys ;
       iter (extract defs sref cycle) depends
     with Not_found -> ()
 
@@ -526,7 +526,7 @@ let bind sigma defs xs =
 
 let get_class sigma xs x =
   List.sort Var.compare
-    (List.filter (fun y -> Vars.mem y xs) (Sigma.class_of sigma x))
+    (List.filter (fun y -> Vars.mem y xs) (Subst.class_of sigma x))
 
 let rec add_eq ps y = function
   | z::zs -> add_eq (p_equal (e_var y) (e_var z) :: ps) y zs
@@ -540,7 +540,7 @@ let add_definitions sigma defs xs ps =
   Vars.fold
     (fun x ps ->
        let ps = add_equals (get_class sigma xs x) ps in
-       try F.p_equal (e_var x) (Sigma.find x sigma) :: ps
+       try F.p_equal (e_var x) (Subst.find x sigma) :: ps
        with Not_found -> ps
     ) xs ps
 
diff --git a/src/plugins/wp/Letify.mli b/src/plugins/wp/Letify.mli
index 08a30c9c0370d3d70b2108b1c423ad3c8c3a1e35..105b063affb36d2c1416537b4f8fab12966d7b43 100644
--- a/src/plugins/wp/Letify.mli
+++ b/src/plugins/wp/Letify.mli
@@ -42,7 +42,7 @@ sig
 
 end
 
-module Sigma :
+module Subst :
 sig
   type t
   val equal : t -> t -> bool
@@ -68,12 +68,12 @@ sig
   val domain : t -> Vars.t
 end
 
-val bind : Sigma.t -> Defs.t -> Vars.t -> Sigma.t
+val bind : Subst.t -> Defs.t -> Vars.t -> Subst.t
 (** [bind sigma defs xs] select definitions in [defs]
     targeting variables [xs]. The result is a new substitution that
     potentially augment [sigma] with definitions for [xs] (and others). *)
 
-val add_definitions : Sigma.t -> Defs.t -> Vars.t -> pred list -> pred list
+val add_definitions : Subst.t -> Defs.t -> Vars.t -> pred list -> pred list
 (** [add_definitions sigma defs xs ps] keep all
     definitions of variables [xs] from [sigma] that comes from [defs].
     They are added to [ps]. *)
diff --git a/src/plugins/wp/LogicAssigns.ml b/src/plugins/wp/LogicAssigns.ml
index 374a114c4df9274d85945e6eb4aee1d36b9517f2..4d45ae61bb541c0846efdb2eeeb831aae7f51357 100644
--- a/src/plugins/wp/LogicAssigns.ml
+++ b/src/plugins/wp/LogicAssigns.ml
@@ -20,17 +20,16 @@
 (*                                                                        *)
 (**************************************************************************)
 
-open Sigs
+open Memory
 
 module Make
-    ( M : Sigs.Model )
-    ( L : Sigs.LogicSemantics with module M = M ) =
+    ( M : Memory.Model )
+    ( L : Memory.LogicSemantics with module M = M ) =
 struct
 
   module M = M
   module L = L
-  open M
-  module D = Heap.Set
+  module D = Sigma.Domain
 
   (* -------------------------------------------------------------------------- *)
   (* --- Domain                                                             --- *)
@@ -40,7 +39,7 @@ struct
     | Sloc l | Sdescr(_,l,_) -> M.domain obj l
     | Srange(l,obj,_,_) | Sarray(l,obj,_) -> M.domain obj l
 
-  let domain (r: loc Sigs.region) =
+  let domain (r: L.loc Memory.region) =
     List.fold_left
       (fun d (obj,sloc) -> D.union d (dsloc obj sloc)) D.empty r
 
@@ -49,23 +48,23 @@ struct
   (* -------------------------------------------------------------------------- *)
 
   let rec assigned_seq hs s = function
-    | [] -> Bag.concat (M.Sigma.assigned ~pre:s.pre ~post:s.post D.empty) hs
+    | [] -> Bag.concat (Sigma.assigned ~pre:s.pre ~post:s.post D.empty) hs
 
     | [obj,sloc] ->
       let eq_sloc = M.assigned s obj sloc in
       let hs_sloc = Bag.list (List.map Cvalues.equation eq_sloc) in
-      let hs_sdom = M.Sigma.assigned ~pre:s.pre ~post:s.post (dsloc obj sloc) in
+      let hs_sdom = Sigma.assigned ~pre:s.pre ~post:s.post (dsloc obj sloc) in
       Bag.concat (Bag.concat hs_sloc hs_sdom) hs
 
     | (obj,sloc)::tail ->
-      let sigma = M.Sigma.havoc s.post (dsloc obj sloc) in
+      let sigma = Sigma.havoc s.post (dsloc obj sloc) in
       let s_local = { pre = sigma ; post = s.post } in
       let s_other = { pre = s.pre ; post = sigma } in
       let eq_sloc = M.assigned s_local obj sloc in
       let hs_sloc = Bag.list (List.map Cvalues.equation eq_sloc) in
       assigned_seq (Bag.concat hs_sloc hs) s_other tail
 
-  let apply_assigns (s:sigma sequence) (r: M.loc Sigs.region) =
+  let apply_assigns (s: Sigma.sigma sequence) (r: M.loc Memory.region) =
     Bag.elements (assigned_seq Bag.empty s r)
 
 end
diff --git a/src/plugins/wp/LogicAssigns.mli b/src/plugins/wp/LogicAssigns.mli
index 8c33ad6b6a3f4aabcb02e9f996d5c95dc0990668..f5b4225f1d5b119575a6b8909410886919a9a4b7 100644
--- a/src/plugins/wp/LogicAssigns.mli
+++ b/src/plugins/wp/LogicAssigns.mli
@@ -21,6 +21,6 @@
 (**************************************************************************)
 
 module Make
-    ( M : Sigs.Model )
-    ( L : Sigs.LogicSemantics with module M = M ) :
-  Sigs.LogicAssigns with module M = M and module L = L
+    ( M : Memory.Model )
+    ( L : Memory.LogicSemantics with module M = M ) :
+  Memory.LogicAssigns with module M = M and module L = L
diff --git a/src/plugins/wp/LogicCompiler.ml b/src/plugins/wp/LogicCompiler.ml
index 6af277cb4ee8acf30c4a56e68b24f492c08f7a11..93af6967881d936e641554271b23be6d0fbf50b3 100644
--- a/src/plugins/wp/LogicCompiler.ml
+++ b/src/plugins/wp/LogicCompiler.ml
@@ -32,14 +32,15 @@ open Clabels
 open Ctypes
 open Lang
 open Lang.F
-open Sigs
+open Memory
+open Sigma
 open Definitions
 
 let dkey_lemma = Wp_parameters.register_category "lemma"
 
 type polarity = [ `Positive | `Negative | `NoPolarity ]
 
-module Make( M : Sigs.Model ) =
+module Make( M : Memory.Model ) =
 struct
 
   (* -------------------------------------------------------------------------- *)
@@ -48,11 +49,9 @@ struct
 
   open M
 
-  type value = M.loc Sigs.value
-  type logic = M.loc Sigs.logic
-  type result = loc Sigs.result
-  type sigma = M.Sigma.t
-  type chunk = M.Chunk.t
+  type value = M.loc Memory.value
+  type logic = M.loc Memory.logic
+  type result = loc Memory.result
 
   type signature =
     | CST of Integer.t
@@ -129,7 +128,7 @@ struct
   type call = {
     kf : kernel_function ;
     formals : value Varinfo.Map.t ;
-    mutable result : M.loc Sigs.result option ;
+    mutable result : M.loc Memory.result option ;
     mutable status : var option ;
   }
 
@@ -267,7 +266,7 @@ struct
     assert (not (Clabels.is_here label));
     try LabelMap.find label frame.labels
     with Not_found ->
-      let s = M.Sigma.create () in
+      let s = Sigma.create () in
       frame.labels <- LabelMap.add label s frame.labels ; s
 
   let set_at_frame frame label sigma =
@@ -390,13 +389,13 @@ struct
       let heap = List.fold_left
           (fun m x ->
              let obj = object_of x.vtype in
-             M.Sigma.Chunk.Set.union m (M.domain obj (M.cvar x))
-          ) M.Sigma.Chunk.Set.empty vars
+             Domain.union m (M.domain obj (M.cvar x))
+          ) Domain.empty vars
       in List.fold_left
         (fun acc l ->
            let label = Clabels.of_logic l in
            let sigma = Sigma.create () in
-           M.Sigma.Chunk.Set.fold_sorted
+           Domain.fold_sorted
              (fun chunk (parm,sigm) ->
                 let x = Sigma.get sigma chunk in
                 let s = Sig_chunk (chunk,label) in
@@ -455,7 +454,7 @@ struct
         let (parm,sigm) =
           LabelMap.fold
             (fun label sigma acc ->
-               M.Sigma.Chunk.Set.fold_sorted
+               Domain.fold_sorted
                  (fun chunk acc ->
                     if filter result (Sigma.get sigma chunk) then
                       let (parm,sigm) = acc in
@@ -476,7 +475,7 @@ struct
   let cc_logic : (env -> Cil_types.term -> logic) ref
     = ref (fun _ _ -> assert false)
   let cc_region
-    : (env -> Cil_types.term -> loc Sigs.region) ref
+    : (env -> Cil_types.term -> loc Memory.region) ref
     = ref (fun _ -> assert false)
 
   let term env t = !cc_term env t
@@ -707,10 +706,10 @@ struct
       let frame = logic_frame l.l_var_info.lv_name l.l_tparams in
       in_frame frame
         (fun () ->
-           Heap.Map.fold_sorted
+           Sigma.Heap.Map.fold_sorted
              (fun chunk labels acc ->
-                let basename = Chunk.basename_of_chunk chunk in
-                let tau = Chunk.tau_of_chunk chunk in
+                let basename = Sigma.Chunk.basename_of_chunk chunk in
+                let tau = Sigma.Chunk.tau_of_chunk chunk in
                 LabelSet.fold
                   (fun label (parm,sigm) ->
                      let x = Lang.freshvar ~basename tau in
@@ -882,7 +881,7 @@ struct
   (* --- Binding Formal with Actual w.r.t Signature                         --- *)
   (* -------------------------------------------------------------------------- *)
 
-  let rec bind_labels env phi_labels labels : M.Sigma.t LabelMap.t =
+  let rec bind_labels env phi_labels labels : sigma LabelMap.t =
     match phi_labels, labels with
     | [], [] -> LabelMap.empty
     | l1 :: phi_labels, l2 :: labels ->
@@ -909,7 +908,7 @@ struct
             with Not_found ->
               Wp_parameters.fatal "*** Label %a not-found@." Clabels.pretty l
           in
-          M.Sigma.value sigma c
+          Sigma.value sigma c
       ) sparam
 
   let call_fun env
diff --git a/src/plugins/wp/LogicCompiler.mli b/src/plugins/wp/LogicCompiler.mli
index e662c2e3e592fa8275d86902ceaa6247fc487c23..aca6f9fbac4331ab716196f4db1c823cee41b26f 100644
--- a/src/plugins/wp/LogicCompiler.mli
+++ b/src/plugins/wp/LogicCompiler.mli
@@ -28,23 +28,21 @@ open LogicUsage
 open Cil_types
 open Cil_datatype
 open Clabels
-open Lang
 open Lang.F
-open Sigs
+open Memory
+open Sigma
 open Definitions
 
 type polarity = [ `Positive | `Negative | `NoPolarity ]
 
-module Make( M : Sigs.Model ) :
+module Make( M : Memory.Model ) :
 sig
 
   (** {3 Definitions} *)
 
-  type value = M.loc Sigs.value
-  type logic = M.loc Sigs.logic
-  type result = M.loc Sigs.result
-  type sigma = M.Sigma.t
-  type chunk = M.Chunk.t
+  type value = M.loc Memory.value
+  type logic = M.loc Memory.logic
+  type result = M.loc Memory.result
 
   (** {3 Frames} *)
 
@@ -101,12 +99,12 @@ sig
   val term : env -> Cil_types.term -> term
   val pred : polarity -> env -> predicate -> pred
   val logic : env -> Cil_types.term -> logic
-  val region : env -> Cil_types.term -> M.loc Sigs.region
+  val region : env -> Cil_types.term -> M.loc Memory.region
 
   val bootstrap_term : (env -> Cil_types.term -> term) -> unit
   val bootstrap_pred : (polarity -> env -> predicate -> pred) -> unit
   val bootstrap_logic : (env -> Cil_types.term -> logic) -> unit
-  val bootstrap_region : (env -> Cil_types.term -> M.loc Sigs.region) -> unit
+  val bootstrap_region : (env -> Cil_types.term -> M.loc Memory.region) -> unit
 
   (** {3 Application} *)
 
diff --git a/src/plugins/wp/LogicSemantics.ml b/src/plugins/wp/LogicSemantics.ml
index 94825bb06bb83b2f42558fcf50ad0e5561702770..b51c72c389c166a0223d8567697300765053ead9 100644
--- a/src/plugins/wp/LogicSemantics.ml
+++ b/src/plugins/wp/LogicSemantics.ml
@@ -34,21 +34,18 @@ open Ctypes
 open Lang
 open Lang.F
 open Definitions
-open Sigs
+open Memory
 
-module Make(M : Sigs.Model) =
+module Make(M : Memory.Model) =
 struct
 
-  module M = M
-  open M
-
   type loc = M.loc
-  type value = loc Sigs.value
-  type logic = loc Sigs.logic
-  type result = loc Sigs.result
-  type region = loc Sigs.region
-  type sigma = Sigma.t
+  type value = loc Memory.value
+  type logic = loc Memory.logic
+  type result = loc Memory.result
+  type region = loc Memory.region
 
+  module M = M
   module L = Cvalues.Logic(M)
   module C = LogicCompiler.Make(M)
 
@@ -200,9 +197,9 @@ struct
     match lhost with
     | TResult ty ->
       begin match C.result () with
-        | Sigs.R_var x ->
+        | Memory.R_var x ->
           access_offset env (Vexp (e_var x)) loffset
-        | Sigs.R_loc l ->
+        | Memory.R_loc l ->
           load_loc env ty l loffset
       end
     | TMem e ->
diff --git a/src/plugins/wp/LogicSemantics.mli b/src/plugins/wp/LogicSemantics.mli
index 63379500ac3dd6dbdfaf8b9937be99aeb5539baa..785c8feebf5f3e2eaf1d29e7b838fd8fa4f8d1ba 100644
--- a/src/plugins/wp/LogicSemantics.mli
+++ b/src/plugins/wp/LogicSemantics.mli
@@ -24,4 +24,4 @@
 (* --- ACSL Translation                                                   --- *)
 (* -------------------------------------------------------------------------- *)
 
-module Make(M : Sigs.Model) : Sigs.LogicSemantics with module M = M
+module Make(M : Memory.Model) : Memory.LogicSemantics with module M = M
diff --git a/src/plugins/wp/MemAddr.ml b/src/plugins/wp/MemAddr.ml
index 8305c9438f828f7ec1f8af4b15f8b9741b46a9f4..79529ee65807410f6cebfe6328e84630220f8edf 100644
--- a/src/plugins/wp/MemAddr.ml
+++ b/src/plugins/wp/MemAddr.ml
@@ -384,19 +384,19 @@ type range =
   | RANGE of term * Vset.set (* base - range offset *)
 
 let range ~shift ~addrof ~sizeof = function
-  | Sigs.Rloc(obj,loc) ->
+  | Memory.Rloc(obj,loc) ->
     LOC( addrof loc , sizeof obj )
-  | Sigs.Rrange(loc,obj,Some a,Some b) ->
+  | Memory.Rrange(loc,obj,Some a,Some b) ->
     let s = sizeof obj in
     let p = addrof (shift loc obj a) in
     let n = e_mul s (e_range a b) in
     LOC( p , n )
-  | Sigs.Rrange(loc,_obj,None,None) ->
+  | Memory.Rrange(loc,_obj,None,None) ->
     RANGE( base (addrof loc) , Vset.range None None )
-  | Sigs.Rrange(loc,obj,Some a,None) ->
+  | Memory.Rrange(loc,obj,Some a,None) ->
     let s = sizeof obj in
     RANGE( base (addrof loc) , Vset.range (Some (e_mul s a)) None )
-  | Sigs.Rrange(loc,obj,None,Some b) ->
+  | Memory.Rrange(loc,obj,None,Some b) ->
     let s = sizeof obj in
     RANGE( base (addrof loc) , Vset.range None (Some (e_mul s b)) )
 
diff --git a/src/plugins/wp/MemAddr.mli b/src/plugins/wp/MemAddr.mli
index 85ffa9bfa8a282d9e7fc1061ba40844f0ca581f4..5bdc87da7a0fe74bc3c03a7703078c0cb280c505 100644
--- a/src/plugins/wp/MemAddr.mli
+++ b/src/plugins/wp/MemAddr.mli
@@ -152,7 +152,7 @@ val included :
   shift:('loc -> Ctypes.c_object -> term -> 'loc) ->
   addrof:('loc -> term) ->
   sizeof:(Ctypes.c_object -> term) ->
-  'loc Sigs.rloc -> 'loc Sigs.rloc -> pred
+  'loc Memory.rloc -> 'loc Memory.rloc -> pred
 (** [included ~shift ~addrof ~sizeof r1 r2] builds a predicate that checks
     whether [r1] is included in [r2].
     - [shift loc obj k]: [loc] shifted of [k] [obj] in the memory model,
@@ -164,7 +164,7 @@ val separated :
   shift:('loc -> Ctypes.c_object -> term -> 'loc) ->
   addrof:('loc -> term) ->
   sizeof:(Ctypes.c_object -> term) ->
-  'loc Sigs.rloc -> 'loc Sigs.rloc -> pred
+  'loc Memory.rloc -> 'loc Memory.rloc -> pred
 (** [separated ~shift ~addrof ~sizeof r1 r2] builds a predicate that checks
     whether [r1] and [r2] are separated.
     - [shift loc obj k]: [loc] shifted of [k] [obj] in the memory model,
diff --git a/src/plugins/wp/MemBytes.ml b/src/plugins/wp/MemBytes.ml
index eb593e762bfc2efa349c5f801f96e5f2b7214962..2dc4d78d107c837918b7a512094bfcf0eb9eb9aa 100644
--- a/src/plugins/wp/MemBytes.ml
+++ b/src/plugins/wp/MemBytes.ml
@@ -21,7 +21,7 @@
 (**************************************************************************)
 
 open Lang.F
-open Sigs
+open Memory
 open Ctypes
 
 module Logic = Qed.Logic
@@ -126,7 +126,6 @@ end
 (* Model *)
 let datatype = "MemBytes"
 let lc_name = String.lowercase_ascii datatype
-let dkey_state = Wp_parameters.register_category (lc_name ^ ":state")
 let dkey_model = Wp_parameters.register_category (lc_name ^ ":model")
 
 let configure () =
@@ -181,23 +180,22 @@ struct
     | Init -> "init"
     | Alloc -> "alloc"
 
+  let is_init = function Init -> true | Mem | Alloc -> false
+  let is_primary _ = false
   let is_framed _ = false
 end
 
-module Heap = Qed.Collection.Make(Chunk)
-module Sigma = Sigma.Make(Chunk)(Heap)
+module State = Sigma.Make(Chunk)
 
-type loc = term
+let m_alloc = State.chunk Alloc
+let m_init = State.chunk Init
+let m_mem = State.chunk Mem
 
-let pretty fmt =
-  Format.fprintf fmt "l:(%a)" pp_term
+type loc = term
 
 let vars = vars
 let occurs = occurs
 
-type chunk = Chunk.t
-type sigma = Sigma.t
-type domain = Sigma.domain
 type segment = loc rloc
 
 let shift_cluster () =
@@ -382,20 +380,15 @@ module Shift = WpContext.Generator(Cobj)
         dfun.d_lfun
     end)
 
-let field loc f =
-  Wp_parameters.debug ~level:3 ~dkey:dkey_model
-    "%s.field %a.%a" datatype pretty loc Cil_printer.pp_field f ;
-  e_fun (ShiftField.get f) [loc]
-let shift loc obj k =
-  Wp_parameters.debug ~level:3 ~dkey:dkey_model
-    "%s.shift %a + %a(%a)" datatype pretty loc pp_term k Ctypes.pp_object obj ;
-  e_fun (Shift.get obj) [loc;k]
+let field loc f = e_fun (ShiftField.get f) [loc]
+let shift loc obj k = e_fun (Shift.get obj) [loc;k]
 
 (* ********************************************************************** *)
 (* VALIDITY and SEPARATION                                                *)
 (* ********************************************************************** *)
 
-let allocated sigma l = e_get (Sigma.value sigma Alloc) (MemAddr.base l)
+let allocated sigma l =
+  e_get (Sigma.value sigma m_alloc) (MemAddr.base l)
 
 let s_valid sigma acs p n =
   let valid = match acs with
@@ -403,10 +396,10 @@ let s_valid sigma acs p n =
     | RD -> MemAddr.valid_rd
     | OBJ -> (fun m p _ -> MemAddr.valid_obj m p)
   in
-  valid (Sigma.value sigma Alloc) p n
+  valid (Sigma.value sigma m_alloc) p n
 
 let s_invalid sigma p n =
-  MemAddr.invalid (Sigma.value sigma Alloc) p n
+  MemAddr.invalid (Sigma.value sigma m_alloc) p n
 
 let segment phi = function
   | Rloc(obj,l) ->
@@ -559,7 +552,7 @@ let load_int_raw memory kind addr =
   read memory addr
 
 let load_int sigma kind addr =
-  load_int_raw (Sigma.value sigma Chunk.Mem) kind addr
+  load_int_raw (Sigma.value sigma m_mem) kind addr
 
 let load_float sigma kind addr =
   int_to_float kind @@ load_int sigma (Float.ikind kind) addr
@@ -579,7 +572,7 @@ let load_init memory size loc =
   | _ -> assert false
 
 let is_init_atom sigma obj loc =
-  let init_memory = Sigma.value sigma Init in
+  let init_memory = Sigma.value sigma m_init in
   let size = sizeof_object obj in
   load_init init_memory size loc
 
@@ -595,7 +588,7 @@ let store_int sigma kind addr v =
     | UInt64 -> WBytes.write_uint64
     | SInt64 -> WBytes.write_sint64
   in
-  Chunk.Mem, write (Sigma.value sigma Mem) addr v
+  m_mem, write (Sigma.value sigma m_mem) addr v
 
 let store_float sigma kind addr v =
   store_int sigma (Float.ikind kind) addr @@ float_to_int kind v
@@ -614,13 +607,11 @@ let store_init_raw m size loc v =
   write m loc v
 
 let set_init_atom sigma obj loc v =
-  let init_memory = Sigma.value sigma Init in
+  let init_memory = Sigma.value sigma m_init in
   let size = sizeof_object obj in
-  Chunk.Init, store_init_raw init_memory size loc v
+  m_init, store_init_raw init_memory size loc v
 
 module Model = struct
-  module Chunk = Chunk
-  module Sigma = Sigma
 
   let name = "MemBytes.Loader"
 
@@ -634,8 +625,8 @@ module Model = struct
   let to_region_pointer l = 0,l
   let of_region_pointer _r _obj l = l
 
-  let value_footprint _ _ = Sigma.Chunk.Set.singleton Chunk.Mem
-  let init_footprint _ _ = Sigma.Chunk.Set.singleton Chunk.Init
+  let value_footprint _ _ = Sigma.Domain.singleton m_mem
+  let init_footprint _ _ = Sigma.Domain.singleton m_init
 
   let frames  ~addr:p ~offset:n ?(basename="mem") tau =
     let t_block = Qed.Logic.Array (Qed.Logic.Int, tau) in
@@ -657,23 +648,27 @@ module Model = struct
       "havoc"  , []    , [sep]     , m , mh ;
     ]
 
-  let frames obj addr = function
-    | Chunk.Alloc -> []
-    | m ->
+  let frames obj addr chunk =
+    match Sigma.ckind chunk with
+    | State.Mu Alloc -> []
+    | State.Mu m ->
       let offset = sizeof obj in
       let tau = Chunk.val_of_chunk m in
       let basename = Chunk.basename_of_chunk m in
       frames ~addr ~offset ~basename tau
+    | _ -> []
 
   let last sigma obj l =
     let n = protected_sizeof_object obj in
     e_sub (e_div (allocated sigma l) n) e_one
 
   let memcpy obj ~mtgt ~msrc ~ltgt ~lsrc ~length chunk =
-    if chunk <> Chunk.Alloc then
+    match Sigma.ckind chunk with
+    | State.Mu Alloc -> msrc
+    | State.Mu _ ->
       let n = e_mul (e_int @@ sizeof_object obj) length in
       WBytes.memcpy mtgt msrc ltgt lsrc n
-    else msrc
+    | _ -> assert false
 
   let eqmem_forall obj loc _chunk m1 m2 =
     let xp = Lang.freshvar ~basename:"p" MemAddr.t_addr in
@@ -697,7 +692,7 @@ module Model = struct
   let is_init_atom = is_init_atom
   let is_init_range sigma obj loc length =
     let n = e_mul (sizeof obj) length in
-    WBytes.is_init_range (Sigma.value sigma Init) loc n
+    WBytes.is_init_range (Sigma.value sigma m_init) loc n
 
   let set_init_atom = set_init_atom
   let set_init obj loc ~length _chunk ~current =
@@ -872,7 +867,7 @@ module STRING = WpContext.Generator(LITERAL)
         let i = Lang.freshvar ~basename:"i" Lang.t_int in
         let c = Cstring.char_at cst (e_var i) in
         let ikind = Ctypes.c_char () in
-        let m = Lang.freshvar ~basename:"mchar" (Chunk.tau_of_chunk Mem) in
+        let m = Lang.freshvar ~basename:"mchar" @@ Chunk.tau_of_chunk Mem in
         let addr = shift (MemAddr.global base) (C_int ikind) (e_var i) in
         let v = load_int_raw (e_var m) ikind addr in
         let read = Lang.F.(p_equal c v) in
@@ -936,8 +931,6 @@ let global _sigma p =
 (* STATE                                                                  *)
 (* ********************************************************************** *)
 
-type state = chunk Tmap.t
-
 let rec lookup_a e =
   match repr e with
   | Fun( f , [e] ) when MemAddr.is_f_global f -> lookup_a e
@@ -950,45 +943,24 @@ and lookup_f f es =
     | RS_Index _ , [e;k] -> Mstate.index (lookup_lv e) k
     | _ -> raise Not_found
   with Not_found when es = [] ->
-    Sigs.(Mvar (RegisterBASE.find f),[])
-
-and lookup_lv e = try lookup_a e with Not_found -> Sigs.(Mmem e,[])
+    Memory.(Mvar (RegisterBASE.find f),[])
 
-let mchunk c =
-  match c with
-  | Chunk.Init -> Sigs.Mchunk (Pretty_utils.to_string Chunk.pretty c, KInit)
-  | _ -> Sigs.Mchunk (Pretty_utils.to_string Chunk.pretty c, KValue)
+and lookup_lv e = try lookup_a e with Not_found -> Memory.(Mmem e,[])
 
 let lookup s e =
-  Wp_parameters.debug ~level:2 ~dkey:dkey_state "%s.lookup _ %a"
-    datatype pp_term e ;
-  try mchunk (Tmap.find e s)
-  with Not_found ->
-  try match repr e with
-    | Fun( f , es ) -> Sigs.Maddr (lookup_f f es)
-    | Aget( m , k ) when Tmap.find m s = Init ->
-      Sigs.Mlval (lookup_lv k, KInit)
-    | Aget( m , k ) when Tmap.find m s <> Alloc ->
-      Sigs.Mlval (lookup_lv k, KValue)
-    | _ -> Sigs.Mterm
-  with Not_found -> Sigs.Mterm
-
-let state sigma =
-  Wp_parameters.debug ~level:2 ~dkey:dkey_state "%s.state _" datatype ;
-  let s = ref Tmap.empty in
-  Sigma.iter (fun c x -> s := Tmap.add (e_var x) c !s) sigma ; !s
-
-let iter f s =
-  Wp_parameters.debug ~level:2 ~dkey:dkey_state "%s.iter _ _" datatype ;
-  Tmap.iter (fun m c -> f (mchunk c) m) s
-
-let updates _ _ =
-  Wp_parameters.debug ~level:2 ~dkey:dkey_state "%s.updates _ _" datatype ;
-  Bag.empty
-
-let apply f s =
-  Wp_parameters.debug ~level:2 ~dkey:dkey_state "%s.apply _ _" datatype ;
-  Tmap.fold (fun m c w -> Tmap.add (f m) c w) s Tmap.empty
+  match repr e with
+  | Fun( f , es ) -> Memory.Maddr (lookup_f f es)
+  | Aget( m , k ) ->
+    begin
+      match Sigma.ckind @@ Tmap.find m s with
+      | State.Mu Alloc -> raise Not_found
+      | State.Mu Init -> Memory.Minit (lookup_lv k)
+      | State.Mu _ -> Memory.Mlval (lookup_lv k)
+      | _ -> raise Not_found
+    end
+  | _ -> raise Not_found
+
+let updates _ _ = Bag.empty
 
 (* ********************************************************************** *)
 (* POINTERS OPS                                                           *)
@@ -1008,7 +980,7 @@ let base_offset loc =
 let block_length sigma _obj loc =
   Wp_parameters.debug ~level:3 ~dkey:dkey_model
     "%s.block_length _ _ _ " datatype ;
-  e_get (Sigma.value sigma Chunk.Alloc) (MemAddr.base loc)
+  e_get (Sigma.value sigma m_alloc) (MemAddr.base loc)
 
 let cast _ loc =
   Wp_parameters.debug ~level:3 ~dkey:dkey_model
@@ -1028,7 +1000,7 @@ let int_of_loc _ loc =
 
 (* -------------------------------------------------------------------------- *)
 
-let domain _ _ = Sigma.Chunk.Set.of_list [ Init ; Mem ]
+let domain _ _ = Sigma.Domain.of_list [ m_init ; m_mem ]
 
 let is_null = p_equal null
 let loc_eq = p_equal
@@ -1090,14 +1062,14 @@ let frame sigma =
     then [ phi (Sigma.value sigma chunk) ]
     else []
   in
-  wellformed_frame MemAddr.linked Alloc @
-  wellformed_frame WBytes.cinits Init @
-  wellformed_frame WBytes.sconst Mem @
-  [ framed (Sigma.value sigma Mem) ]
+  wellformed_frame MemAddr.linked m_alloc @
+  wellformed_frame WBytes.cinits m_init @
+  wellformed_frame WBytes.sconst m_mem @
+  [ framed (Sigma.value sigma m_mem) ]
 
 let is_well_formed s =
   Wp_parameters.debug ~level:2 ~dkey:dkey_model "%s.is_well_formed _" datatype ;
-  WBytes.bytes (Sigma.value s Mem)
+  WBytes.bytes (Sigma.value s m_mem)
 
 (* ********************************************************************** *)
 (* ALLOCATION                                                             *)
@@ -1107,25 +1079,25 @@ let alloc sigma xs =
   Wp_parameters.debug ~level:2 ~dkey:dkey_model
     "%s.alloc %a %a"
     datatype Sigma.pretty sigma (Pretty_utils.pp_list Cil_printer.pp_varinfo) xs ;
-  if xs = [] then sigma else Sigma.havoc_chunk sigma Alloc
+  if xs = [] then sigma else Sigma.havoc_chunk sigma m_alloc
 
 let scope seq scope xs =
   Wp_parameters.debug ~level:2 ~dkey:dkey_model
     "%s.scope { %a ; %a } %s %a"
     datatype Sigma.pretty seq.pre Sigma.pretty seq.post
-    (if scope = Sigs.Enter then "Enter" else "Leave")
+    (if scope = Memory.Enter then "Enter" else "Leave")
     (Pretty_utils.pp_list Cil_printer.pp_varinfo) xs ;
   if xs = [] then [] else
     let alloc =
       List.fold_left
         (fun m x ->
            let size = match scope with
-             | Sigs.Leave -> e_zero
-             | Sigs.Enter ->
+             | Memory.Leave -> e_zero
+             | Memory.Enter ->
                protected_sizeof_object @@ Ctypes.object_of x.Cil_types.vtype
            in e_set m (BASE.get x) size)
-        (Sigma.value seq.pre Alloc) xs in
-    [ p_equal (Sigma.value seq.post Alloc) alloc ]
+        (Sigma.value seq.pre m_alloc) xs in
+    [ p_equal (Sigma.value seq.post m_alloc) alloc ]
 
 (* ********************************************************************** *)
 (* API with Region                                                        *)
diff --git a/src/plugins/wp/MemDebug.ml b/src/plugins/wp/MemDebug.ml
index 85923c2dec570b403b3d3f0e6a9ebeb27777f8de..403cd95c3bad0db1f775e41c1240a2889f5bdaf2 100644
--- a/src/plugins/wp/MemDebug.ml
+++ b/src/plugins/wp/MemDebug.ml
@@ -22,7 +22,7 @@
 
 open Cil_datatype
 open Lang.F
-open Sigs
+open Memory
 
 (* ------------------------------------------------------------------------ *)
 (* ----  Pretty-printers                                               ---- *)
@@ -79,23 +79,15 @@ let debug_access = Wp_parameters.debug ~dkey:dkey_access
 let debug_valid = Wp_parameters.debug ~dkey:dkey_valid
 let debug_alias = Wp_parameters.debug ~dkey:dkey_alias
 
-module Make(M : Sigs.Model) =
+module Make(M : Memory.Model) =
 struct
   let datatype = "MemDebug." ^ M.datatype
   let configure = M.configure
 
   let hypotheses = M.hypotheses
 
-  module Chunk = M.Chunk
-
-  module Heap = M.Heap
-  module Sigma = M.Sigma
-
   type loc = M.loc
-  type chunk = M.chunk
-  type sigma = M.sigma
   type segment = M.segment
-  type state = M.state
 
   (* ---------------------------------------------------------------------- *)
   (* ----  Pretty-printing                                             ---- *)
@@ -103,11 +95,8 @@ struct
 
   let pretty = M.pretty
 
-  let state = M.state
-  let iter = M.iter
   let lookup = M.lookup
   let updates = M.updates
-  let apply = M.apply
 
   let vars = M.vars
   let occurs = M.occurs
@@ -273,7 +262,6 @@ struct
   let frame = M.frame
   let is_well_formed = M.is_well_formed
   let base_offset = M.base_offset
-  type domain = M.domain
   let configure_ia = M.configure_ia
 
 end
diff --git a/src/plugins/wp/MemDebug.mli b/src/plugins/wp/MemDebug.mli
index 08c054eb2fe439e36541bb8c5f0031e5facfc8ef..70ffcd5c0a30edd79432284ae06c68717a72e9a0 100644
--- a/src/plugins/wp/MemDebug.mli
+++ b/src/plugins/wp/MemDebug.mli
@@ -25,14 +25,14 @@
 (* -------------------------------------------------------------------------- *)
 
 val pp_sequence : 'a Pretty_utils.formatter -> Format.formatter ->
-  'a Sigs.sequence -> unit
-val pp_equation : Format.formatter -> Sigs.equation -> unit
-val pp_acs : Format.formatter -> Sigs.acs -> unit
+  'a Memory.sequence -> unit
+val pp_equation : Format.formatter -> Memory.equation -> unit
+val pp_acs : Format.formatter -> Memory.acs -> unit
 val pp_value : 'a Pretty_utils.formatter -> Format.formatter ->
-  'a Sigs.value -> unit
+  'a Memory.value -> unit
 val pp_rloc : 'a Pretty_utils.formatter -> Format.formatter ->
-  'a Sigs.rloc -> unit
+  'a Memory.rloc -> unit
 val pp_sloc : 'a Pretty_utils.formatter -> Format.formatter ->
-  'a Sigs.sloc -> unit
+  'a Memory.sloc -> unit
 
-module Make(_ : Sigs.Model) : Sigs.Model
+module Make(_ : Memory.Model) : Memory.Model
diff --git a/src/plugins/wp/MemEmpty.ml b/src/plugins/wp/MemEmpty.ml
index 0cdd629bd9ce63d95fef1c8e973b5fabfcd6a159..4df6d13091ab269e8f18dbe0d90d4a67de2dfbc1 100644
--- a/src/plugins/wp/MemEmpty.ml
+++ b/src/plugins/wp/MemEmpty.ml
@@ -25,7 +25,7 @@
 (* -------------------------------------------------------------------------- *)
 
 open Lang.F
-open Sigs
+open Memory
 
 module Logic = Qed.Logic
 
@@ -45,33 +45,10 @@ let configure_ia _ = no_binder
 
 let hypotheses p = p
 
-module Chunk =
-struct
-  type t = unit
-  let self = "empty"
-  let hash () = 0
-  let equal () () = true
-  let compare () () = 0
-  let pretty _fmt () = ()
-  let tau_of_chunk () = Logic.Int
-  let basename_of_chunk () = "u"
-  let is_framed () = true
-end
-
-module Heap = Qed.Collection.Make(Chunk)
-module Sigma = Sigma.Make(Chunk)(Heap)
-
 type loc = unit
-type chunk = Chunk.t
-type sigma = Sigma.t
-type domain = Sigma.domain
 type segment = loc rloc
-type state = unit
-let state _ = ()
-let iter _ _ = ()
-let lookup _ _ = Mterm
+let lookup _ _ = raise Not_found
 let updates _ _ = Bag.empty
-let apply _ _ = ()
 
 let pretty _fmt () = ()
 let vars _l = Vars.empty
@@ -93,7 +70,7 @@ let cast _ _l = ()
 let loc_of_int _ _ = ()
 let int_of_loc _ () = e_zero
 
-let domain _obj _l = Sigma.Chunk.Set.empty
+let domain _obj _l = Sigma.Domain.empty
 let is_well_formed _s = p_true
 
 let source = "Empty Model"
diff --git a/src/plugins/wp/MemEmpty.mli b/src/plugins/wp/MemEmpty.mli
index 8523b1df8cafd65934b46f4554c4c8e850896068..4fc1b752cb1eee50883707fc41a1604fa6e3ea75 100644
--- a/src/plugins/wp/MemEmpty.mli
+++ b/src/plugins/wp/MemEmpty.mli
@@ -24,4 +24,4 @@
 (* --- Empty Memory Model                                                 --- *)
 (* -------------------------------------------------------------------------- *)
 
-include Sigs.Model
+include Memory.Model
diff --git a/src/plugins/wp/MemLoader.ml b/src/plugins/wp/MemLoader.ml
index 8e8d64305041afb57db6c87a13e17c9bf2ed7cba..98ccad0f9162bb44fedfeb4045fa973ff2614014 100644
--- a/src/plugins/wp/MemLoader.ml
+++ b/src/plugins/wp/MemLoader.ml
@@ -30,7 +30,8 @@ open Definitions
 open Ctypes
 open Lang
 open Lang.F
-open Sigs
+open Memory
+open Sigma
 
 (* -------------------------------------------------------------------------- *)
 (* --- Compound Loader                                                    --- *)
@@ -42,9 +43,6 @@ let cluster () =
 module type Model =
 sig
 
-  module Chunk : Chunk
-  module Sigma : Sigma with type chunk = Chunk.t
-
   val name : string
 
   type loc
@@ -56,12 +54,12 @@ sig
   val to_region_pointer : loc -> int * term
   val of_region_pointer : int -> c_object -> term -> loc
 
-  val value_footprint: c_object -> loc -> Sigma.domain
-  val init_footprint: c_object -> loc -> Sigma.domain
+  val value_footprint: c_object -> loc -> domain
+  val init_footprint: c_object -> loc -> domain
 
   val frames : c_object -> loc -> Chunk.t -> frame list
 
-  val last : Sigma.t -> c_object -> loc -> term
+  val last : sigma -> c_object -> loc -> term
 
   val memcpy : c_object -> mtgt:term -> msrc:term -> ltgt:loc -> lsrc:loc ->
     length:term -> Chunk.t -> term
@@ -69,17 +67,17 @@ sig
   val eqmem_forall :
     c_object -> loc -> Chunk.t -> term -> term -> var list * pred * pred
 
-  val load_int : Sigma.t -> c_int -> loc -> term
-  val load_float : Sigma.t -> c_float -> loc -> term
-  val load_pointer : Sigma.t -> typ -> loc -> loc
+  val load_int : sigma -> c_int -> loc -> term
+  val load_float : sigma -> c_float -> loc -> term
+  val load_pointer : sigma -> typ -> loc -> loc
 
-  val store_int : Sigma.t -> c_int -> loc -> term -> Chunk.t * term
-  val store_float : Sigma.t -> c_float -> loc -> term -> Chunk.t * term
-  val store_pointer : Sigma.t -> typ -> loc -> term -> Chunk.t * term
+  val store_int : sigma -> c_int -> loc -> term -> Chunk.t * term
+  val store_float : sigma -> c_float -> loc -> term -> Chunk.t * term
+  val store_pointer : sigma -> typ -> loc -> term -> Chunk.t * term
 
-  val is_init_atom : Sigma.t -> c_object -> loc -> term
-  val is_init_range : Sigma.t -> c_object -> loc -> term -> pred
-  val set_init_atom : Sigma.t -> c_object -> loc -> term -> Chunk.t * term
+  val is_init_atom : sigma -> c_object -> loc -> term
+  val is_init_range : sigma -> c_object -> loc -> term -> pred
+  val set_init_atom : sigma -> c_object -> loc -> term -> Chunk.t * term
   val set_init : c_object -> loc -> length:term ->
     Chunk.t -> current:term -> term
 
@@ -88,12 +86,6 @@ end
 module Make (M : Model) =
 struct
 
-  type chunk = M.Chunk.t
-
-  module Chunk = M.Chunk
-  module Sigma = M.Sigma
-  module Domain = M.Sigma.Chunk.Set
-
   let signature ft =
     let s = Sigma.create () in
     let xs = ref [] in
@@ -106,7 +98,7 @@ struct
     List.rev !xs , List.rev !cs , s
 
   let domain obj loc =
-    M.Sigma.Chunk.Set.union
+    Domain.union
       (M.value_footprint obj loc)
       (M.init_footprint obj loc)
 
@@ -209,12 +201,12 @@ struct
 
   module type LOAD_INFO = sig
     val kind : Lang.datakind
-    val footprint : c_object -> M.loc -> M.Sigma.domain
+    val footprint : c_object -> M.loc -> domain
     val t_comp : compinfo -> Lang.tau
     val t_array : AKEY.base -> Lang.tau
     val comp_id : compinfo -> string
     val array_id : AKEY.base -> string
-    val load : M.Sigma.t -> c_object -> M.loc -> term
+    val load : sigma -> c_object -> M.loc -> term
   end
 
   module VALUE_LOAD_INFO = struct
@@ -373,9 +365,9 @@ struct
 
   module LOADER_GEN
       (ATOM: sig
-         val load_int : M.Sigma.t -> c_int -> M.loc -> term
-         val load_float : M.Sigma.t -> c_float -> M.loc -> term
-         val load_pointer : M.Sigma.t -> typ -> M.loc -> term
+         val load_int : sigma -> c_int -> M.loc -> term
+         val load_float : sigma -> c_float -> M.loc -> term
+         val load_pointer : sigma -> typ -> M.loc -> term
        end)
       (COMP: sig val get : (int*compinfo) -> (lfun * chunk list) end)
       (ARRAY: sig val get : (int*AKEY.base*Matrix.t) -> (lfun * chunk list) end)
@@ -418,7 +410,7 @@ struct
   let () = VALUE_LOAD_INFO.load_rec := load_value
 
   let load sigma obj loc =
-    let open Sigs in
+    let open Memory in
     match obj with
     | C_int i -> Val (M.load_int sigma i loc)
     | C_float f -> Val (M.load_float sigma f loc)
diff --git a/src/plugins/wp/MemLoader.mli b/src/plugins/wp/MemLoader.mli
index 18a8b7acf484fad80dcd30e418545a38db161593..dad5319318004c70e7d182ccb53a90f2d24bf0f5 100644
--- a/src/plugins/wp/MemLoader.mli
+++ b/src/plugins/wp/MemLoader.mli
@@ -28,7 +28,8 @@ open Cil_types
 open Definitions
 open Ctypes
 open Lang.F
-open Sigs
+open Memory
+open Sigma
 
 val cluster : unit -> cluster
 
@@ -36,9 +37,6 @@ val cluster : unit -> cluster
 module type Model =
 sig
 
-  module Chunk : Chunk
-  module Sigma : Sigma with type chunk = Chunk.t
-
   val name : string
 
   type loc
@@ -52,12 +50,12 @@ sig
   val to_region_pointer : loc -> int * term
   val of_region_pointer : int -> c_object -> term -> loc
 
-  val value_footprint: c_object -> loc -> Sigma.domain
-  val init_footprint: c_object -> loc -> Sigma.domain
+  val value_footprint: c_object -> loc -> domain
+  val init_footprint: c_object -> loc -> domain
 
-  val frames : c_object -> loc -> Chunk.t -> frame list
+  val frames : c_object -> loc -> chunk -> frame list
 
-  val last : Sigma.t -> c_object -> loc -> term
+  val last : sigma -> c_object -> loc -> term
 
   val memcpy : c_object -> mtgt:term -> msrc:term -> ltgt:loc -> lsrc:loc ->
     length:term -> Chunk.t -> term
@@ -65,17 +63,17 @@ sig
   val eqmem_forall :
     c_object -> loc -> Chunk.t -> term -> term -> var list * pred * pred
 
-  val load_int : Sigma.t -> c_int -> loc -> term
-  val load_float : Sigma.t -> c_float -> loc -> term
-  val load_pointer : Sigma.t -> typ -> loc -> loc
+  val load_int : sigma -> c_int -> loc -> term
+  val load_float : sigma -> c_float -> loc -> term
+  val load_pointer : sigma -> typ -> loc -> loc
 
-  val store_int : Sigma.t -> c_int -> loc -> term -> Chunk.t * term
-  val store_float : Sigma.t -> c_float -> loc -> term -> Chunk.t * term
-  val store_pointer : Sigma.t -> typ -> loc -> term -> Chunk.t * term
+  val store_int : sigma -> c_int -> loc -> term -> Chunk.t * term
+  val store_float : sigma -> c_float -> loc -> term -> Chunk.t * term
+  val store_pointer : sigma -> typ -> loc -> term -> Chunk.t * term
 
-  val is_init_atom : Sigma.t -> c_object -> loc -> term
-  val is_init_range : Sigma.t -> c_object -> loc -> term -> pred
-  val set_init_atom : Sigma.t -> c_object -> loc -> term -> Chunk.t * term
+  val is_init_atom : sigma -> c_object -> loc -> term
+  val is_init_range : sigma -> c_object -> loc -> term -> pred
+  val set_init_atom : sigma -> c_object -> loc -> term -> Chunk.t * term
   val set_init : c_object -> loc -> length:term ->
     Chunk.t -> current:term -> term
 
@@ -85,23 +83,23 @@ end
 module Make (M : Model) :
 sig
 
-  val domain : c_object -> M.loc -> M.Sigma.domain
+  val domain : c_object -> M.loc -> domain
 
-  val load : M.Sigma.t -> c_object -> M.loc -> M.loc Sigs.value
-  val load_init : M.Sigma.t -> c_object -> M.loc -> term
-  val load_value : M.Sigma.t -> c_object -> M.loc -> term
+  val load : sigma -> c_object -> M.loc -> M.loc Memory.value
+  val load_init : sigma -> c_object -> M.loc -> term
+  val load_value : sigma -> c_object -> M.loc -> term
 
-  val memcpy : M.Sigma.t sequence -> c_object -> ?lsrc:M.loc -> M.loc -> equation list
-  val memcpy_length : M.Sigma.t sequence -> c_object -> ?lsrc:M.loc -> M.loc -> term -> equation list
+  val memcpy : sigma sequence -> c_object -> ?lsrc:M.loc -> M.loc -> equation list
+  val memcpy_length : sigma sequence -> c_object -> ?lsrc:M.loc -> M.loc -> term -> equation list
 
-  val stored : M.Sigma.t sequence -> c_object -> M.loc -> term -> equation list
-  val stored_init : M.Sigma.t sequence -> c_object -> M.loc -> term -> equation list
-  val copied : M.Sigma.t sequence -> c_object -> M.loc -> M.loc -> equation list
-  val copied_init : M.Sigma.t sequence -> c_object -> M.loc -> M.loc -> equation list
+  val stored : sigma sequence -> c_object -> M.loc -> term -> equation list
+  val stored_init : sigma sequence -> c_object -> M.loc -> term -> equation list
+  val copied : sigma sequence -> c_object -> M.loc -> M.loc -> equation list
+  val copied_init : sigma sequence -> c_object -> M.loc -> M.loc -> equation list
 
-  val assigned : M.Sigma.t sequence -> c_object -> M.loc sloc -> equation list
+  val assigned : sigma sequence -> c_object -> M.loc sloc -> equation list
 
-  val initialized : M.Sigma.t -> M.loc rloc -> pred
+  val initialized : sigma -> M.loc rloc -> pred
 
 end
 
diff --git a/src/plugins/wp/MemMemory.mli b/src/plugins/wp/MemMemory.mli
index ed0cd5ecd87acc9e490e56461627065187d30d90..de1dd94c5d214e342196f9346df502611d19ec7b 100644
--- a/src/plugins/wp/MemMemory.mli
+++ b/src/plugins/wp/MemMemory.mli
@@ -61,7 +61,7 @@ val framed : term -> pred
 *)
 
 val frames : addr:term -> offset:term -> sizeof:term ->
-  ?basename:string -> tau -> Sigs.frame list
+  ?basename:string -> tau -> Memory.frame list
 
 (** {2 Unsupported Union Fields} *)
 
diff --git a/src/plugins/wp/MemRegion.ml b/src/plugins/wp/MemRegion.ml
index 24d8175a3ac8f07e7e48c73d1e3c3e199e12edaf..a979eb3edef1efdc6f949597a26aa2acb290eb4f 100644
--- a/src/plugins/wp/MemRegion.ml
+++ b/src/plugins/wp/MemRegion.ml
@@ -27,11 +27,12 @@
 open Cil_types
 open Ctypes
 open Lang.F
-open Sigs
+open Memory
+open Sigma
 open MemMemory
 
-type primitive = Int of c_int | Float of c_float | Ptr
-type kind = Single of primitive | Many of primitive | Garbled
+type prim = Int of c_int | Float of c_float | Ptr
+type kind = Single of prim | Many of prim | Garbled
 
 let pp_prim fmt = function
   | Int i -> Ctypes.pp_int fmt i
@@ -55,9 +56,9 @@ let tau_of_prim = function
 module type RegionProxy =
 sig
   type region
-  module Type : Sigs.Type with type t = region
   val id : region -> int
   val of_id : int -> region option
+  val pretty : Format.formatter -> region -> unit
   val kind : region -> kind
   val name : region -> string option
   val cvar : varinfo -> region option
@@ -75,8 +76,7 @@ end
 (* -------------------------------------------------------------------------- *)
 
 module type ModelWithLoader = sig
-  include Sigs.Model
-  include Sigs.Model
+  include Memory.Model
   val sizeof : c_object -> term
 
   val last : sigma -> c_object -> loc -> term
@@ -108,7 +108,7 @@ end
 (* --- Region Memory Model                                                --- *)
 (* -------------------------------------------------------------------------- *)
 
-module Make (R:RegionProxy) (M:ModelWithLoader) (*: Sigs.Model*) =
+module Make (R:RegionProxy) (M:ModelWithLoader) (*: Memory.Model*) =
 struct
 
   type region = R.region
@@ -119,200 +119,60 @@ struct
   let configure_ia = M.configure_ia
   let hypotheses = M.hypotheses
 
-  module MChunk = M.Chunk
-  module RChunk =
+  module Chunk =
   struct
-    let self = "MemRegion.RChunk"
+    let self = "MemRegion.Chunk"
 
-    type mu = Value of primitive | Array of primitive | ValInit | ArrInit
-    type t = { mu : mu ; region : R.region }
+    type data = Value of prim | Array of prim | ValInit | ArrInit
+    type t = { data : data ; region : R.region }
 
-    let pp_mu fmt = function
+    let pp_data fmt = function
       | Value p -> Format.fprintf fmt "µ%a" pp_prim p
       | Array p -> Format.fprintf fmt "µ%a[]" pp_prim p
       | ValInit -> Format.pp_print_string fmt "µinit"
       | ArrInit -> Format.pp_print_string fmt "µinit[]"
 
-    let hash { mu ; region } = Hashtbl.hash (mu, R.Type.hash region)
-    let equal a b = Stdlib.(=) a.mu b.mu && R.Type.equal a.region b.region
+    let hash { data ; region } = Hashtbl.hash (data, R.id region)
+    let equal a b = Stdlib.(=) a.data b.data && R.id a.region = R.id b.region
     let compare a b =
-      let cmp = Stdlib.compare a.mu b.mu in
-      if cmp <> 0 then cmp else R.Type.compare a.region b.region
+      let cmp = Stdlib.compare a.data b.data in
+      if cmp <> 0 then cmp else Int.compare (R.id a.region) (R.id b.region)
 
-    let pretty fmt { mu ; region } =
-      Format.fprintf fmt "%a@%03d" pp_mu mu (R.id region)
+    let pretty fmt { data ; region } =
+      Format.fprintf fmt "%a@%03d" pp_data data (R.id region)
 
-    let tau_of_chunk { mu } =
-      match mu with
+    let tau_of_chunk { data } =
+      match data with
       | Value p -> tau_of_prim p
       | ValInit -> Qed.Logic.Bool
       | Array p -> Qed.Logic.Array(MemAddr.t_addr,tau_of_prim p)
       | ArrInit -> Qed.Logic.Array(MemAddr.t_addr,Qed.Logic.Bool)
 
-    let basename_of_chunk { mu ; region } =
-      match mu with
+    let basename_of_chunk c =
+      match c.data with
       | ValInit -> "Vinit"
       | ArrInit -> "Minit"
       | Array p -> Format.asprintf "M%a" pp_prim p
       | Value p ->
-        match R.name region with
+        match R.name c.region with
         | Some a -> a
         | None -> Format.asprintf "V%a" pp_prim p
 
-    let is_framed _ = false
-
-  end
-
-  module MHeap = M.Heap
-  module MSigma = M.Sigma
-
-  module RHeap = Qed.Collection.Make(RChunk)
-  module RSigma = Sigma.Make(RChunk)(RHeap)
-
-  type chunk =
-    | M of MChunk.t
-    | R of RChunk.t
+    let is_init c =
+      match c.data with
+      | ValInit | ArrInit -> true
+      | Array _ | Value _ -> false
 
-  let cmap f g (m,r) = (f m, g r)
-  let cmap2 f g (m1,r1) (m2,r2) = (f m1 m2, g r1 r2)
-  let capply f g (m,r) = function M c -> f m c, r | R c -> m, g r c
-  let cmerge f g (m,r) = function M c -> f m c | R c -> g r c
-  let mseq { pre ; post } = { pre = fst pre ; post = fst post }
-  let rseq { pre ; post } = { pre = snd pre ; post = snd post }
+    let is_primary c =
+      match c.data with
+      | Value _ -> true
+      | ValInit | ArrInit | Array _ -> false
 
-  module Chunk : Sigs.Chunk with type t = chunk =
-  struct
-    type t = chunk
-    let self = "Wp.MemRegion.Self"
-    let hash = function
-      | M c -> 3 * MChunk.hash c
-      | R c -> 5 * RChunk.hash c
-
-    let equal ca cb = match ca, cb with
-      | M c1, M c2 -> MChunk.equal c1 c2
-      | R c1, R c2 -> RChunk.equal c1 c2
-      | M _, R _ | R _, M _ -> false
-
-    let compare c1 c2 =
-      match c1, c2 with
-      | M m1, M m2 -> MChunk.compare m1 m2
-      | R r1, R r2 -> RChunk.compare r1 r2
-      | M _, R _ -> (-1)
-      | R _, M _ -> (+1)
-
-    let pretty fmt = function
-      | M c -> MChunk.pretty fmt c
-      | R c -> RChunk.pretty fmt c
-
-    let tau_of_chunk = function
-      | M c -> MChunk.tau_of_chunk c
-      | R c -> RChunk.tau_of_chunk c
-
-    let basename_of_chunk = function
-      | M c -> MChunk.basename_of_chunk c
-      | R c -> RChunk.basename_of_chunk c
-
-    let is_framed = function
-      | M c -> MChunk.is_framed c
-      | R c -> RChunk.is_framed c
+    let is_framed _ = false
 
   end
 
-  module Heap = Qed.Collection.Make(Chunk)
-  module Domain = Heap.Set
-
-  type sigma = MSigma.t * RSigma.t
-  type domain = Domain.t
-
-  module Sigma =
-  struct
-
-    type t = sigma
-    type chunk = Chunk.t
-    module Chunk = Heap
-
-    type domain = Domain.t
-
-    let create () : t = (MSigma.create (), RSigma.create ())
-
-    let pretty fmt (m,r) =
-      Format.fprintf fmt "@[<hv 0>{@[<hv 2>@ %a;@ %a;@]@ }@]"
-        MSigma.pretty m
-        RSigma.pretty r
-
-    let empty : domain = Domain.empty
-    let union = Domain.union
-    let mem = cmerge MSigma.mem RSigma.mem
-    let get = cmerge MSigma.get RSigma.get
-    let value = cmerge MSigma.value RSigma.value
-    let copy = cmap MSigma.copy RSigma.copy
-    let choose = cmap2 MSigma.choose RSigma.choose
-    let havoc_chunk = capply MSigma.havoc_chunk RSigma.havoc_chunk
-    let havoc_any ~call = cmap (MSigma.havoc_any ~call) (RSigma.havoc_any ~call)
-
-    let merge (m1,r1) (m2,r2) =
-      let m,p1,p2 = MSigma.merge m1 m2 in
-      let r,q1,q2 = RSigma.merge r1 r2 in
-      (m,r), Passive.union p1 q1, Passive.union p2 q2
-
-    let merge_list l =
-      let m,p = MSigma.merge_list @@ List.map fst l in
-      let r,q = RSigma.merge_list @@ List.map snd l in
-      (m,r), List.map2 Passive.union p q
-
-    let join (m1,r1) (m2,r2) =
-      Passive.union (MSigma.join m1 m2) (RSigma.join r1 r2)
-
-    let iter f (m,r) =
-      begin
-        MSigma.iter (fun c -> f (M c)) m ;
-        RSigma.iter (fun c -> f (R c)) r ;
-      end
-
-    let iter2 f (m1,r1) (m2,r2) =
-      begin
-        MSigma.iter2 (fun c -> f (M c)) m1 m2 ;
-        RSigma.iter2 (fun c -> f (R c)) r1 r2 ;
-      end
-
-    let mdomain d =
-      MHeap.Set.fold (fun c s -> Domain.add (M c) s) d Domain.empty
-    let rdomain d =
-      RHeap.Set.fold (fun c s -> Domain.add (R c) s) d Domain.empty
-
-    let dsplit d =
-      let m = ref MHeap.Set.empty in
-      let r = ref RHeap.Set.empty in
-      Domain.iter
-        (function
-          | M c -> m := MHeap.Set.add c !m
-          | R c -> r := RHeap.Set.add c !r
-        ) d ;
-      !m, !r
-
-    let assigned ~pre:(m1,r1) ~post:(m2,r2) d =
-      let m,r = dsplit d in
-      let hm = MSigma.assigned ~pre:m1 ~post:m2 m in
-      let hr = RSigma.assigned ~pre:r1 ~post:r2 r in
-      Bag.concat hm hr
-
-    let havoc (m,r) d =
-      let dm,dr = dsplit d in
-      MSigma.havoc m dm, RSigma.havoc r dr
-
-    let remove_chunks (m,r) d =
-      let dm,dr = dsplit d in
-      MSigma.remove_chunks m dm, RSigma.remove_chunks r dr
-
-    let domain (m,r) =
-      union (mdomain @@ MSigma.domain m) (rdomain @@ RSigma.domain r)
-
-    let writes (s : sigma sequence) =
-      union
-        (mdomain @@ MSigma.writes @@ mseq s)
-        (rdomain @@ RSigma.writes @@ rseq s)
-
-  end
+  module State = Sigma.Make(Chunk)
 
   (* -------------------------------------------------------------------------- *)
   (* --- Region Loader                                                         --- *)
@@ -322,9 +182,6 @@ struct
   struct
     let name = "MemRegion.Loader"
 
-    module Chunk = Chunk
-    module Sigma = Sigma
-
     type loc =
       | Null
       | Raw of M.loc
@@ -340,7 +197,7 @@ struct
     (* --- Utilities on locations                                         --- *)
     (* ---------------------------------------------------------------------- *)
 
-    let last sigma ty l = M.last (fst sigma) ty (loc l)
+    let last sigma ty l = M.last sigma ty (loc l)
 
     let to_addr l = M.pointer_val (loc l)
 
@@ -357,35 +214,39 @@ struct
       | Single Ptr | Many Ptr | Garbled ->
         let offset = M.sizeof ty in
         let sizeof = Lang.F.e_one in
-        let tau = Chunk.tau_of_chunk c in
-        let basename = Chunk.basename_of_chunk c in
+        let tau = Sigma.Chunk.tau_of_chunk c in
+        let basename = Sigma.Chunk.basename_of_chunk c in
         MemMemory.frames ~addr:(to_addr l) ~offset ~sizeof ~basename tau
       | _ -> []
 
     let memcpy ty ~mtgt ~msrc ~ltgt ~lsrc ~length chunk =
-      match chunk with
-      | M c ->
-        M.memcpy ty ~mtgt ~msrc ~ltgt:(loc ltgt) ~lsrc:(loc lsrc) ~length c
-      | R c ->
-        match c.mu with
-        | Value _ | ValInit -> msrc
-        | Array _ | ArrInit ->
-          e_fun f_memcpy [mtgt;msrc;to_addr ltgt;to_addr lsrc;length]
+      match Sigma.ckind chunk with
+      | State.Mu { data } ->
+        begin
+          match data with
+          | Value _ | ValInit -> msrc
+          | Array _ | ArrInit ->
+            e_fun f_memcpy [mtgt;msrc;to_addr ltgt;to_addr lsrc;length]
+        end
+      | _ ->
+        M.memcpy ty ~mtgt ~msrc ~ltgt:(loc ltgt) ~lsrc:(loc lsrc) ~length chunk
 
     let eqmem_forall ty l chunk m1 m2 =
-      match chunk with
-      | M c -> M.eqmem_forall ty (loc l) c m1 m2
-      | R c ->
-        match c.mu with
-        | Value _ | ValInit -> [], p_true, p_equal m1 m2
-        | Array _ | ArrInit ->
-          let xp = Lang.freshvar ~basename:"b" MemAddr.t_addr in
-          let p = e_var xp in
-          let n = M.sizeof ty in
-          let separated =
-            p_call MemAddr.p_separated [p;e_one;to_addr l;n] in
-          let equal = p_equal (e_get m1 p) (e_get m2 p) in
-          [xp],separated,equal
+      match Sigma.ckind chunk with
+      | State.Mu { data } ->
+        begin
+          match data with
+          | Value _ | ValInit -> [], p_true, p_equal m1 m2
+          | Array _ | ArrInit ->
+            let xp = Lang.freshvar ~basename:"b" MemAddr.t_addr in
+            let p = e_var xp in
+            let n = M.sizeof ty in
+            let separated =
+              p_call MemAddr.p_separated [p;e_one;to_addr l;n] in
+            let equal = p_equal (e_get m1 p) (e_get m2 p) in
+            [xp],separated,equal
+        end
+      | _ -> M.eqmem_forall ty (loc l) chunk m1 m2
 
     (* ---------------------------------------------------------------------- *)
     (* --- Load                                                           --- *)
@@ -406,7 +267,7 @@ struct
     let of_region_pointer r _ t =
       make (M.pointer_loc t) (R.of_id r)
 
-    let check_access action (p:primitive) (q:primitive) =
+    let check_access action (p:prim) (q:prim) =
       if Stdlib.(<>) p q then
         Warning.error ~source:"MemRegion"
           "Inconsistent %s (%a <> %a)"
@@ -415,27 +276,27 @@ struct
     let load_int sigma iota loc : term =
       let l,r = localized "load int" loc in
       match R.kind r with
-      | Garbled -> M.load_int (fst sigma) iota l
+      | Garbled -> M.load_int sigma iota l
       | Single p ->
         check_access "load" p (Int iota) ;
-        RSigma.value (snd sigma) { mu = Value p ; region = r }
+        State.value sigma { data = Value p ; region = r }
       | Many p ->
         check_access "load" p (Int iota) ;
         e_get
-          (RSigma.value (snd sigma) { mu = Array p ; region = r})
+          (State.value sigma { data = Array p ; region = r})
           (M.pointer_val l)
 
     let load_float sigma flt loc : term =
       let l,r = localized "load float" loc in
       match R.kind r with
-      | Garbled -> M.load_float (fst sigma) flt l
+      | Garbled -> M.load_float sigma flt l
       | Single p ->
         check_access "load" p (Float flt) ;
-        RSigma.value (snd sigma) { mu = Value p ; region = r }
+        State.value sigma { data = Value p ; region = r }
       | Many p ->
         check_access "load" p (Float flt) ;
         e_get
-          (RSigma.value (snd sigma) { mu = Array p ; region = r})
+          (State.value sigma { data = Array p ; region = r})
           (M.pointer_val l)
 
     let load_pointer sigma ty loc : loc =
@@ -445,20 +306,20 @@ struct
         Warning.error ~source:"MemRegion"
           "Attempt to load pointer without points-to@\n\
            (addr %a, region %a)"
-          M.pretty l R.Type.pretty r
+          M.pretty l R.pretty r
       | Some _ as rp ->
         let loc =
           match R.kind r with
-          | Garbled -> M.load_pointer (fst sigma) ty l
+          | Garbled -> M.load_pointer sigma ty l
           | Single p ->
             check_access "load" p Ptr ;
             M.pointer_loc @@
-            RSigma.value (snd sigma) { mu = Value p ; region = r }
+            State.value sigma { data = Value p ; region = r }
           | Many p ->
             check_access "load" p Ptr ;
             M.pointer_loc @@
             e_get
-              (RSigma.value (snd sigma) { mu = Array p ; region = r})
+              (State.value sigma { data = Array p ; region = r})
               (M.pointer_val l)
         in make loc rp
 
@@ -466,44 +327,44 @@ struct
     (* --- Store                                                          --- *)
     (* ---------------------------------------------------------------------- *)
 
-    let store_int sigma iota loc v : Chunk.t * term =
+    let store_int sigma iota loc v : Sigma.chunk * term =
       let l,r = localized "store int" loc in
       match R.kind r with
       | Garbled ->
-        let c, m = M.store_int (fst sigma) iota l v in M c, m
+        M.store_int sigma iota l v
       | Single p ->
         check_access "store" p (Int iota) ;
-        R { mu = Value p ; region = r }, v
+        State.chunk { data = Value p ; region = r }, v
       | Many p ->
         check_access "store" p (Int iota) ;
-        let rc = RChunk.{ mu = Array p ; region = r } in
-        R rc, e_set (RSigma.value (snd sigma) rc) (M.pointer_val l) v
+        let rc = Chunk.{ data = Array p ; region = r } in
+        State.chunk rc, e_set (State.value sigma rc) (M.pointer_val l) v
 
-    let store_float sigma flt loc v : Chunk.t * term =
+    let store_float sigma flt loc v : Sigma.chunk * term =
       let l,r = localized "store float" loc in
       match R.kind r with
       | Garbled ->
-        let c,m = M.store_float (fst sigma) flt l v in M c, m
+        M.store_float sigma flt l v
       | Single p ->
         check_access "store" p (Float flt) ;
-        R { mu = Value p ; region = r }, v
+        State.chunk { data = Value p ; region = r }, v
       | Many p ->
         check_access "store" p (Float flt) ;
-        let rc = RChunk.{ mu = Array p ; region = r } in
-        R rc, e_set (RSigma.value (snd sigma) rc) (M.pointer_val l) v
+        let rc = Chunk.{ data = Array p ; region = r } in
+        State.chunk rc, e_set (State.value sigma rc) (M.pointer_val l) v
 
-    let store_pointer sigma ty loc v : Chunk.t * term =
+    let store_pointer sigma ty loc v : Sigma.chunk * term =
       let l,r = localized "store pointer" loc in
       match R.kind r with
       | Garbled ->
-        let c,m = M.store_pointer (fst sigma) ty l v in M c, m
+        M.store_pointer sigma ty l v
       | Single p ->
         check_access "store" p Ptr ;
-        R { mu = Value p ; region = r }, v
+        State.chunk { data = Value p ; region = r }, v
       | Many p ->
         check_access "store" p Ptr ;
-        let rc = RChunk.{ mu = Array p ; region = r } in
-        R rc, e_set (RSigma.value (snd sigma) rc) (M.pointer_val l) v
+        let rc = Chunk.{ data = Array p ; region = r } in
+        State.chunk rc, e_set (State.value sigma rc) (M.pointer_val l) v
 
     (* ---------------------------------------------------------------------- *)
     (* --- Init                                                           --- *)
@@ -512,42 +373,40 @@ struct
     let is_init_atom sigma ty loc : term =
       let l,r = localized "init atom" loc in
       match R.kind r with
-      | Garbled -> M.is_init_atom (fst sigma) ty l
-      | Single _-> RSigma.value (snd sigma) { mu = ValInit ; region = r }
+      | Garbled -> M.is_init_atom sigma ty l
+      | Single _-> State.value sigma { data = ValInit ; region = r }
       | Many _ ->
         e_get
-          (RSigma.value (snd sigma) { mu = ArrInit ; region = r })
+          (State.value sigma { data = ArrInit ; region = r })
           (M.pointer_val l)
 
-    let set_init_atom sigma ty loc v : Chunk.t * term =
+    let set_init_atom sigma ty loc v : Sigma.chunk * term =
       let l,r = localized "init atom" loc in
       match R.kind r with
       | Garbled ->
-        let c,m = M.set_init_atom (fst sigma) ty l v in M c, m
-      | Single _-> R { mu = ValInit ; region = r }, v
+        M.set_init_atom sigma ty l v
+      | Single _-> State.chunk { data = ValInit ; region = r }, v
       | Many _ ->
-        let rc = RChunk.{ mu = ArrInit ; region = r } in
-        R rc, e_set (RSigma.value (snd sigma) rc) (M.pointer_val l) v
+        let rc = Chunk.{ data = ArrInit ; region = r } in
+        State.chunk rc, e_set (State.value sigma rc) (M.pointer_val l) v
 
     let is_init_range sigma ty loc length : pred =
       let l,r = localized "init atom" loc in
       match R.kind r with
-      | Garbled -> M.is_init_range (fst sigma) ty l length
+      | Garbled -> M.is_init_range sigma ty l length
       | Single _ ->
-        Lang.F.p_bool @@ RSigma.value (snd sigma) { mu = ValInit ; region = r }
+        Lang.F.p_bool @@ State.value sigma { data = ValInit ; region = r }
       | Many _ ->
-        let map = RSigma.value (snd sigma) { mu = ArrInit ; region = r } in
+        let map = State.value sigma { data = ArrInit ; region = r } in
         let size = e_mul (M.sizeof ty) length in
         p_call p_is_init_r [map;M.pointer_val l;size]
 
-
     let set_init ty loc ~length chunk ~current : term =
       let l,r = localized "init atom" loc in
-      match R.kind r, chunk with
-      | Garbled, M c -> M.set_init ty l ~length c ~current
-      | Garbled, R _ -> assert false
-      | Single _, _ -> e_true
-      | Many _ , _ ->
+      match R.kind r with
+      | Garbled -> M.set_init ty l ~length chunk ~current
+      | Single _ -> e_true
+      | Many _ ->
         let size = e_mul (M.sizeof ty) length in
         e_fun f_set_init [current;M.pointer_val l;size]
 
@@ -561,8 +420,8 @@ struct
       else M.init_footprint obj l
 
     let rec footprint ~value obj loc = match loc with
-      | Null  -> Sigma.mdomain @@ mfootprint ~value obj M.null
-      | Raw l -> Sigma.mdomain @@ mfootprint ~value obj l
+      | Null  -> mfootprint ~value obj M.null
+      | Raw l -> mfootprint ~value obj l
       | Loc(l,r) ->
         match obj with
         | C_comp { cfields = None} -> Domain.empty
@@ -578,14 +437,13 @@ struct
           footprint ~value obj (shift loc obj e_zero)
         | C_int _ | C_float _ | C_pointer _ ->
           match R.kind r with
-          | Garbled ->
-            Sigma.mdomain @@ mfootprint ~value obj l
+          | Garbled -> mfootprint ~value obj l
           | Single p ->
-            let mu = if value then RChunk.Value p else ValInit in
-            Sigma.rdomain @@ RHeap.Set.singleton { mu ; region = r }
+            let data = if value then Chunk.Value p else ValInit in
+            State.singleton { data ; region = r }
           | Many p ->
-            let mu = if value then RChunk.Array p else ArrInit in
-            Sigma.rdomain @@ RHeap.Set.singleton { mu ; region = r }
+            let data = if value then Chunk.Array p else ArrInit in
+            State.singleton { data ; region = r }
 
     let value_footprint = footprint ~value:true
     let init_footprint = footprint ~value:false
@@ -595,7 +453,6 @@ struct
   type loc = Loader.loc
   type segment = loc rloc
 
-  open Loader
   module LOADER = MemLoader.Make(Loader)
 
   let load = LOADER.load
@@ -610,64 +467,57 @@ struct
 
   (* {2 Reversing the Model} *)
 
-  type state = M.state
+  let lookup = M.lookup (*TODO: lookups in MemRegion *)
 
-  let state sigma = M.state (fst sigma)
-
-  let lookup s e = M.lookup s e
-
-  let updates = M.updates
-
-  let apply = M.apply
-
-  let iter = M.iter
+  let updates = M.updates (*TODO: updates in MemRegion *)
 
   let pretty fmt (l: loc) =
     match l with
     | Null -> M.pretty fmt M.null
     | Raw l -> M.pretty fmt l
-    | Loc (l,r) -> Format.fprintf fmt "%a@%a" M.pretty l R.Type.pretty r
+    | Loc (l,r) -> Format.fprintf fmt "%a@%a" M.pretty l R.pretty r
 
   (* {2 Memory Model API} *)
 
-  let vars l = M.vars @@ loc l
-  let occurs x l = M.occurs x @@ loc l
-  let null = Null
+  let vars l = M.vars @@ Loader.loc l
+  let occurs x l = M.occurs x @@ Loader.loc l
+  let null = Loader.Null
 
-  let literal ~eid:eid str = make (M.literal ~eid str) (R.literal ~eid str)
+  let literal ~eid:eid str =
+    Loader.make (M.literal ~eid str) (R.literal ~eid str)
 
-  let cvar v = make (M.cvar v) (R.cvar v)
-  let field = field
-  let shift = shift
+  let cvar v = Loader.make (M.cvar v) (R.cvar v)
+  let field = Loader.field
+  let shift = Loader.shift
 
-  let pointer_loc t = Raw (M.pointer_loc t)
-  let pointer_val l = M.pointer_val @@ loc l
-  let base_addr l = Raw (M.base_addr @@ loc l)
-  let base_offset l = M.base_offset @@ loc l
-  let block_length sigma obj l = M.block_length (fst sigma) obj @@ loc l
-  let is_null = function Null -> p_true | Raw l | Loc(l,_) -> M.is_null l
-  let loc_of_int obj t = Raw (M.loc_of_int obj t)
-  let int_of_loc iota l = M.int_of_loc iota @@ loc l
+  let pointer_loc t = Loader.Raw (M.pointer_loc t)
+  let pointer_val l = M.pointer_val @@ Loader.loc l
+  let base_addr l = Loader.Raw (M.base_addr @@ Loader.loc l)
+  let base_offset l = M.base_offset @@ Loader.loc l
+  let block_length sigma obj l = M.block_length sigma obj @@ Loader.loc l
+  let is_null = function Loader.Null -> p_true | Raw l | Loc(l,_) -> M.is_null l
+  let loc_of_int obj t = Loader.Raw (M.loc_of_int obj t)
+  let int_of_loc iota l = M.int_of_loc iota @@ Loader.loc l
 
   let cast conv l =
-    let l0 = loc l in
-    let r0 = reg l in
-    make (M.cast conv l0) r0
+    let l0 = Loader.loc l in
+    let r0 = Loader.reg l in
+    Loader.make (M.cast conv l0) r0
 
-  let loc_eq  a b = M.loc_eq  (loc a) (loc b)
-  let loc_lt  a b = M.loc_lt  (loc a) (loc b)
-  let loc_neq a b = M.loc_neq (loc a) (loc b)
-  let loc_leq a b = M.loc_leq (loc a) (loc b)
-  let loc_diff obj a b = M.loc_diff obj (loc a) (loc b)
+  let loc_eq  a b = M.loc_eq  (Loader.loc a) (Loader.loc b)
+  let loc_lt  a b = M.loc_lt  (Loader.loc a) (Loader.loc b)
+  let loc_neq a b = M.loc_neq (Loader.loc a) (Loader.loc b)
+  let loc_leq a b = M.loc_leq (Loader.loc a) (Loader.loc b)
+  let loc_diff obj a b = M.loc_diff obj (Loader.loc a) (Loader.loc b)
 
   let rloc = function
-    | Rloc(obj, l) -> Rloc (obj, loc l)
-    | Rrange(l, obj, inf, sup) -> Rrange(loc l, obj, inf, sup)
+    | Rloc(obj, l) -> Rloc (obj, Loader.loc l)
+    | Rrange(l, obj, inf, sup) -> Rrange(Loader.loc l, obj, inf, sup)
 
-  let rloc_region = function Rloc(_,l) | Rrange(l,_,_,_) -> reg l
+  let rloc_region = function Rloc(_,l) | Rrange(l,_,_,_) -> Loader.reg l
 
-  let valid sigma acs r = M.valid (fst sigma) acs @@ rloc r
-  let invalid sigma r = M.invalid (fst sigma) (rloc r)
+  let valid sigma acs r = M.valid sigma acs @@ rloc r
+  let invalid sigma r = M.invalid sigma (rloc r)
 
   let included (a : segment) (b : segment) =
     match rloc_region a, rloc_region b with
@@ -679,44 +529,48 @@ struct
     | Some ra, Some rb when R.separated ra rb -> p_true
     | _ -> M.separated (rloc a) (rloc b)
 
-  let alloc sigma vars =
-    if vars = [] then sigma else
-      let m,r = sigma in M.alloc m vars, r
-
-  let scope seq scope vars = M.scope (mseq seq) scope vars
-
-  let global sigma p = M.global (fst sigma) p
+  let alloc = M.alloc
+  let scope = M.scope
+  let global = M.global
 
   let frame sigma =
-    let pool = ref @@ M.frame (fst sigma) in
+    let pool = ref @@ M.frame sigma in
     let assume p = pool := p :: !pool in
-    RSigma.iter
+    Sigma.iter
       (fun c m ->
-         let open RChunk in
-         match c.mu with
-         | ValInit -> ()
-         | ArrInit -> assume @@ MemMemory.cinits (e_var m)
-         | Value Ptr -> assume @@ global sigma (e_var m)
-         | Array Ptr -> assume @@ MemMemory.framed (e_var m)
-         | Value (Int _ | Float _) | Array (Int _ | Float _) -> ()
-      ) (snd sigma) ;
+         match Sigma.ckind c with
+         | State.Mu { data } ->
+           begin
+             match data with
+             | ValInit -> ()
+             | ArrInit -> assume @@ MemMemory.cinits (e_var m)
+             | Value Ptr -> assume @@ global sigma (e_var m)
+             | Array Ptr -> assume @@ MemMemory.framed (e_var m)
+             | Value (Int _ | Float _) | Array (Int _ | Float _) -> ()
+           end
+         | _ -> ()
+      ) sigma ;
     !pool
 
   let is_well_formed sigma =
-    let pool = ref @@ [M.is_well_formed (fst sigma)] in
+    let pool = ref @@ [M.is_well_formed sigma] in
     let assume p = pool := p :: !pool in
-    RSigma.iter
+    Sigma.iter
       (fun c m ->
-         let open RChunk in
-         match c.mu with
-         | ValInit | ArrInit -> ()
-         | Value (Int iota) -> assume @@ Cint.range iota (e_var m)
-         | Array (Int iota) ->
-           let a = Lang.freshvar ~basename:"p" @@ Lang.t_addr () in
-           let b = e_get (e_var m) (e_var a) in
-           assume @@ p_forall [a] (Cint.range iota b)
-         | Value (Float _ | Ptr) | Array (Float _ | Ptr) -> ()
-      ) (snd sigma) ;
+         match Sigma.ckind c with
+         | State.Mu { data } ->
+           begin
+             match data with
+             | ValInit | ArrInit -> ()
+             | Value (Int iota) -> assume @@ Cint.range iota (e_var m)
+             | Array (Int iota) ->
+               let a = Lang.freshvar ~basename:"p" @@ Lang.t_addr () in
+               let b = e_get (e_var m) (e_var a) in
+               assume @@ p_forall [a] (Cint.range iota b)
+             | Value (Float _ | Ptr) | Array (Float _ | Ptr) -> ()
+           end
+         | _ -> ()
+      ) sigma ;
     p_conj !pool
 
 end
diff --git a/src/plugins/wp/MemRegion.mli b/src/plugins/wp/MemRegion.mli
index 985490243680e53c864ec7e9a1fd013314a6c3fa..d3443db2aa4738d946e1529b48ee175b15abd36f 100644
--- a/src/plugins/wp/MemRegion.mli
+++ b/src/plugins/wp/MemRegion.mli
@@ -23,10 +23,11 @@
 open Cil_types
 open Ctypes
 open Lang.F
-open Sigs
+open Memory
+open Sigma
 
-type primitive = | Int of c_int | Float of c_float | Ptr
-type kind = Single of primitive | Many of primitive | Garbled
+type prim = | Int of c_int | Float of c_float | Ptr
+type kind = Single of prim | Many of prim | Garbled
 val pp_kind : Format.formatter -> kind -> unit
 
 (* -------------------------------------------------------------------------- *)
@@ -36,9 +37,9 @@ val pp_kind : Format.formatter -> kind -> unit
 module type RegionProxy =
 sig
   type region
-  module Type : Sigs.Type with type t = region
   val id : region -> int
   val of_id : int -> region option
+  val pretty : Format.formatter -> region -> unit
   val kind : region -> kind
   val name : region -> string option
   val cvar : varinfo -> region option
@@ -53,7 +54,7 @@ end
 
 module type ModelWithLoader =
 sig
-  include Sigs.Model
+  include Memory.Model
   val sizeof : c_object -> term
 
   val last : sigma -> c_object -> loc -> term
@@ -81,4 +82,4 @@ sig
   val init_footprint : c_object -> loc -> domain
 end
 
-module Make : RegionProxy -> ModelWithLoader -> Sigs.Model
+module Make : RegionProxy -> ModelWithLoader -> Memory.Model
diff --git a/src/plugins/wp/MemTyped.ml b/src/plugins/wp/MemTyped.ml
index 1fb70e458188f17617391d164d138b016f64fcc0..bd5012d9745f26cd4f1ae3531101385f0fb6a1c6 100644
--- a/src/plugins/wp/MemTyped.ml
+++ b/src/plugins/wp/MemTyped.ml
@@ -28,8 +28,8 @@ open Cil_types
 open Cil_datatype
 open Ctypes
 open Lang
-open Lang.F
-open Sigs
+open Sigma
+open Memory
 open Definitions
 open MemMemory
 
@@ -46,7 +46,7 @@ let hypotheses p = p
 let configure () =
   begin
     let orig_pointer = Context.push Lang.pointer MemAddr.t_addr in
-    let orig_null    = Context.push Cvalues.null (p_equal MemAddr.null) in
+    let orig_null    = Context.push Cvalues.null (F.p_equal MemAddr.null) in
     let rollback () =
       Context.pop Lang.pointer orig_pointer ;
       Context.pop Cvalues.null orig_null ;
@@ -147,7 +147,7 @@ struct
     | T_alloc -> 12
     | T_init -> 13
   let hash = rank
-  let basename = function
+  let basename_of_chunk = function
     | M_int i when Ctypes.is_char i -> "Mchar"
     | M_int _ -> "Mint"
     | M_f32 -> "Mf32"
@@ -159,7 +159,7 @@ struct
   let equal = (=)
   let pretty fmt = function
     | M_int i -> Format.fprintf fmt "M%a" Ctypes.pp_int i
-    | m -> Format.pp_print_string fmt (basename m)
+    | m -> Format.pp_print_string fmt (basename_of_chunk m)
   let val_of_chunk = function
     | M_int _ -> L.Int
     | M_f32 -> Cfloat.tau_of_float Ctypes.Float32
@@ -174,14 +174,14 @@ struct
     | M_f64 -> L.Array(MemAddr.t_addr,Cfloat.tau_of_float Ctypes.Float64)
     | T_alloc -> t_malloc
     | T_init -> t_init
-  let basename_of_chunk = basename
+  let is_init = function T_init -> true | _ -> false
+  let is_primary _ = false
   let is_framed _ = false
 end
 
-module Heap = Qed.Collection.Make(Chunk)
-module Sigma = Sigma.Make(Chunk)(Heap)
+module State = Sigma.Make(Chunk)
 
-type loc = term (* of type addr *)
+type loc = F.term (* of type addr *)
 
 (* -------------------------------------------------------------------------- *)
 (* --- Utilities on locations                                             --- *)
@@ -190,9 +190,9 @@ type loc = term (* of type addr *)
 let m_float = function Float32 -> M_f32 | Float64 -> M_f64
 
 let rec footprint = function
-  | C_int i -> Heap.Set.singleton (M_int i)
-  | C_float f -> Heap.Set.singleton (m_float f)
-  | C_pointer _ -> Heap.Set.singleton M_pointer
+  | C_int i -> State.singleton (M_int i)
+  | C_float f -> State.singleton (m_float f)
+  | C_pointer _ -> State.singleton M_pointer
   | C_array a -> footprint (object_of a.arr_element)
   | C_comp c -> footprint_comp c
 
@@ -202,18 +202,24 @@ and footprint_comp { cfields } =
   | Some fields ->
     List.fold_left
       (fun ft f ->
-         Heap.Set.union ft (footprint (object_of f.ftype))
-      ) Heap.Set.empty fields
+         Sigma.union ft (footprint (object_of f.ftype))
+      ) Sigma.empty fields
 
 and all_value_chunks () =
-  let ints =
-    List.fold_left (fun l i -> M_int i :: l) []
-      [ Ctypes.CBool ;
-        SInt8 ; UInt8 ; SInt16 ; UInt16 ; SInt32 ; UInt32 ; SInt64 ; UInt64 ]
-  in
-  Heap.Set.of_list (M_pointer :: M_f32 :: M_f64 :: ints)
-
-let init_footprint _ _ = Heap.Set.singleton T_init
+  let pool = ref Domain.empty in
+  let add c = pool := Domain.add (State.chunk c) !pool in
+  let addi i = add (M_int i) in
+  begin
+    add M_pointer ;
+    add M_f32 ; add M_f64 ;
+    addi Ctypes.CBool ;
+    addi SInt8  ; addi UInt8 ;
+    addi SInt16 ; addi UInt16 ;
+    addi SInt32 ; addi UInt32 ;
+    addi SInt64 ; addi UInt64 ;
+  end ; !pool
+
+let init_footprint _ _ = State.singleton T_init
 let value_footprint obj _l = footprint obj
 
 (* Note that it is the length in MemTyped and not the occupied bytes *)
@@ -241,19 +247,19 @@ module OPAQUE_COMP_LENGTH = WpContext.Generator(Cil_datatype.Compinfo)
           l_name = "Positive_Length_of_" ^ Lang.comp_id c ;
           l_triggers = [] ; l_forall = [] ;
           l_cluster = Definitions.compinfo c ;
-          l_lemma = Lang.F.(p_lt e_zero (e_fun size []))
+          l_lemma = F.p_lt F.e_zero (F.e_fun size [])
         } ;
         size
     end)
 
 let rec length_of_object = function
-  | C_int _ | C_float _ | C_pointer _ -> e_one
+  | C_int _ | C_float _ | C_pointer _ -> F.e_one
   | C_comp c -> length_of_comp c
   | C_array { arr_flat = Some { arr_size = n } ; arr_element = elt } ->
-    e_mul (e_int n) (length_of_typ elt)
+    F.e_mul (F.e_int n) (length_of_typ elt)
   | C_array _ as a ->
     if Wp_parameters.ExternArrays.get () then
-      e_int max_int
+      F.e_int max_int
     else
       Warning.error ~source:"Typed Model"
         "Undefined array-size (%a)" Ctypes.pretty a
@@ -263,27 +269,25 @@ and length_of_field f = length_of_typ f.ftype
 and length_of_comp c =
   match c.cfields with
   | None ->
-    Lang.F.e_fun (OPAQUE_COMP_LENGTH.get c) []
+    F.e_fun (OPAQUE_COMP_LENGTH.get c) []
   | Some fields ->
     (* union field are considered as struct field *)
-    e_sum (List.map length_of_field fields)
+    F.e_sum (List.map length_of_field fields)
 
 let position_of_field f =
   let rec fnext k f = function
     | [] -> assert false
     | g::gs ->
       if Fieldinfo.equal f g then k
-      else fnext (e_add k (length_of_field g)) f gs
+      else fnext (F.e_add k (length_of_field g)) f gs
       (* Just as we fail if the field does not exists, we fail
          if we try to get a field position in an opaque struct. *)
-  in fnext e_zero f (Option.get f.fcomp.cfields)
+  in fnext F.e_zero f (Option.get f.fcomp.cfields)
 
 (* -------------------------------------------------------------------------- *)
 (* --- Utilities on loc-as-term                                           --- *)
 (* -------------------------------------------------------------------------- *)
 
-type sigma = Sigma.t
-type domain = Sigma.domain
 type segment = loc rloc
 
 let pretty fmt l = F.pp_term fmt l
@@ -298,19 +302,19 @@ let cluster_globals () =
   Definitions.cluster ~id:"Globals" ~title:"Global Variables" ()
 
 type shift =
-  | RS_Field of fieldinfo * term (* offset of the field *)
-  | RS_Index of term  (* size of the shift *)
+  | RS_Field of fieldinfo * F.term (* offset of the field *)
+  | RS_Index of F.term  (* size of the shift *)
 
 let phi_base = function
   | p::_ -> MemAddr.base p
   | _ -> raise Not_found
 
 let phi_field offset = function
-  | [p] -> e_add (MemAddr.offset p) offset
+  | [p] -> F.e_add (MemAddr.offset p) offset
   | _ -> raise Not_found
 
 let phi_index size = function
-  | [p;k] -> e_add (MemAddr.offset p) (F.e_mul size k)
+  | [p;k] -> F.e_add (MemAddr.offset p) (F.e_mul size k)
   | _ -> raise Not_found
 
 module RegisterShift = WpContext.Static
@@ -333,7 +337,7 @@ module ShiftFieldDef = WpContext.StaticGenerator(Cil_datatype.Fieldinfo)
         let position = position_of_field f in
         (* Since its a generated it is the unique name given *)
         let xloc = Lang.freshvar ~basename:"p" MemAddr.t_addr in
-        let loc = e_var xloc in
+        let loc = F.e_var xloc in
         let def = MemAddr.shift loc position in
         let dfun = Definitions.Function( result , Def , def) in
         RegisterShift.define lfun (RS_Field(f,position)) ;
@@ -394,9 +398,9 @@ module ShiftGen = WpContext.StaticGenerator(Cobj)
         let size = length_of_object obj in
         (* Since its a generated it is the unique name given *)
         let xloc = Lang.freshvar ~basename:"p" MemAddr.t_addr in
-        let loc = e_var xloc in
+        let loc = F.e_var xloc in
         let xk = Lang.freshvar ~basename:"k" Qed.Logic.Int in
-        let k = e_var xk in
+        let k = F.e_var xk in
         let def = MemAddr.shift loc (F.e_mul size k) in
         let dfun = Definitions.Function( result , Def , def) in
         RegisterShift.define shift (RS_Index size) ;
@@ -427,9 +431,9 @@ module Shift = WpContext.Generator(Cobj)
 
 let field l f =
   MemMemory.unsupported_union f ;
-  e_fun (ShiftField.get f) [l]
+  F.e_fun (ShiftField.get f) [l]
 
-let shift l obj k = e_fun (Shift.get obj) [l;k]
+let shift l obj k = F.e_fun (Shift.get obj) [l;k]
 
 module LITERAL =
 struct
@@ -449,12 +453,12 @@ module STRING = WpContext.Generator(LITERAL)
     (struct
       let name = "MemTyped.STRING"
       type key = LITERAL.t
-      type data = term
+      type data = F.term
 
       let linked prefix base cst =
         let name = prefix ^ "_linked" in
         let a = Lang.freshvar ~basename:"alloc" (Chunk.tau_of_chunk T_alloc) in
-        let m = e_var a in
+        let m = F.e_var a in
         let m_linked = MemAddr.linked m in
         let alloc = F.e_get m base in (* The size is alloc-1 *)
         let sized = Cstring.str_len cst (F.e_add alloc F.e_minus_one) in
@@ -462,7 +466,7 @@ module STRING = WpContext.Generator(LITERAL)
           l_kind = Admit ;
           l_name = name ;
           l_triggers = [] ; l_forall = [] ;
-          l_lemma = p_forall [a] (p_imply m_linked sized) ;
+          l_lemma = F.p_forall [a] (F.p_imply m_linked sized) ;
           l_cluster = Cstring.cluster () ;
         }
 
@@ -472,7 +476,7 @@ module STRING = WpContext.Generator(LITERAL)
         Definitions.define_lemma {
           l_kind = Admit ;
           l_name = name ; l_triggers = [] ; l_forall = [] ;
-          l_lemma = p_equal (MemAddr.region base) (e_int re) ;
+          l_lemma = F.p_equal (MemAddr.region base) (F.e_int re) ;
           l_cluster = Cstring.cluster () ;
         }
 
@@ -480,13 +484,13 @@ module STRING = WpContext.Generator(LITERAL)
         (* describe the content of literal strings *)
         let name = prefix ^ "_literal" in
         let i = Lang.freshvar ~basename:"i" L.Int in
-        let c = Cstring.char_at cst (e_var i) in
+        let c = Cstring.char_at cst (F.e_var i) in
         let ikind = Ctypes.c_char () in
-        let addr = shift (MemAddr.global base) (C_int ikind) (e_var i) in
+        let addr = shift (MemAddr.global base) (C_int ikind) (F.e_var i) in
         let m =
           Lang.freshvar ~basename:"mchar" (Chunk.tau_of_chunk (M_int ikind)) in
-        let m_sconst = MemMemory.sconst (e_var m) in
-        let v = F.e_get (e_var m) addr in
+        let m_sconst = MemMemory.sconst (F.e_var m) in
+        let v = F.e_get (F.e_var m) addr in
         let read = F.p_equal c v in
         Definitions.define_lemma {
           l_kind = Admit ;
@@ -541,7 +545,7 @@ module BASE = WpContext.Generator(Varinfo)
     (struct
       let name = "MemTyped.BASE"
       type key = varinfo
-      type data = term
+      type data = F.term
 
       let region prefix x base =
         let name = prefix ^ "_region" in
@@ -549,7 +553,7 @@ module BASE = WpContext.Generator(Varinfo)
         Definitions.define_lemma {
           l_kind = Admit ;
           l_name = name ; l_triggers = [] ; l_forall = [] ;
-          l_lemma = p_equal (MemAddr.region base) (e_int re) ;
+          l_lemma = F.p_equal (MemAddr.region base) (F.e_int re) ;
           l_cluster = cluster_globals () ;
         }
 
@@ -562,19 +566,19 @@ module BASE = WpContext.Generator(Varinfo)
 
       let linked prefix x base =
         let name = prefix ^ "_linked" in
-        let size = if x.vglob then sizeof x else Some e_zero in
+        let size = if x.vglob then sizeof x else Some F.e_zero in
         match size with
         | None -> ()
         | Some size ->
           let a = Lang.freshvar ~basename:"alloc" t_malloc in
-          let m = e_var a in
+          let m = F.e_var a in
           let m_linked = MemAddr.linked m in
-          let base_size = p_equal (F.e_get m base) size in
+          let base_size = F.p_equal (F.e_get m base) size in
           Definitions.define_lemma {
             l_kind = Admit ;
             l_name = name ;
             l_triggers = [] ; l_forall = [] ;
-            l_lemma = p_forall [a] (p_imply m_linked base_size) ;
+            l_lemma = F.p_forall [a] (F.p_imply m_linked base_size) ;
             l_cluster = cluster_globals () ;
           }
 
@@ -582,15 +586,15 @@ module BASE = WpContext.Generator(Varinfo)
         match sizeof x with
         | Some size when Cvalues.always_initialized x ->
           let a = Lang.freshvar ~basename:"init" t_init in
-          let m = e_var a in
+          let m = F.e_var a in
           let init_access =
-            if size = e_one then
-              p_bool (F.e_get m (MemAddr.mk_addr base e_zero))
+            if size = F.e_one then
+              F.p_bool (F.e_get m (MemAddr.mk_addr base F.e_zero))
             else
-              F.p_call p_is_init_r [ m ; MemAddr.mk_addr base e_zero ; size ]
+              F.p_call p_is_init_r [ m ; MemAddr.mk_addr base F.e_zero ; size ]
           in
           let m_init = MemMemory.cinits m in
-          let init_prop = p_forall [a] (p_imply m_init init_access) in
+          let init_prop = F.p_forall [a] (F.p_imply m_init init_access) in
           Definitions.define_lemma {
             l_kind = Admit ;
             l_name = prefix ^ "_init" ;
@@ -612,12 +616,12 @@ module BASE = WpContext.Generator(Varinfo)
         (* Since its a generated it is the unique name given *)
         let prefix = Lang.Fun.debug lfun in
         let vid = if acs_rd then (-x.vid-1) else succ x.vid in
-        let dfun = Definitions.Function( L.Int , Def , e_int vid ) in
+        let dfun = Definitions.Function( L.Int , Def , F.e_int vid ) in
         Definitions.define_symbol {
           d_lfun = lfun ; d_types = 0 ; d_params = [] ; d_definition = dfun ;
           d_cluster = cluster_globals () ;
         } ;
-        let base = e_fun lfun [] in
+        let base = F.e_fun lfun [] in
         RegisterBASE.define lfun x ;
         region prefix x base ;
         linked prefix x base ;
@@ -634,24 +638,24 @@ module BASE = WpContext.Generator(Varinfo)
 let null = MemAddr.null (* as a loc *)
 
 let literal ~eid cst =
-  shift (MemAddr.global (STRING.get (eid,cst))) (C_int (Ctypes.c_char ())) e_zero
+  shift (MemAddr.global (STRING.get (eid,cst))) (C_int (Ctypes.c_char ())) F.e_zero
 
 let cvar x = MemAddr.global (BASE.get x)
 
 let pointer_loc t = t
 let pointer_val t = t
 
-let allocated sigma l = F.e_get (Sigma.value sigma T_alloc) (MemAddr.base l)
+let allocated sigma l = F.e_get (State.value sigma T_alloc) (MemAddr.base l)
 
-let base_addr l = MemAddr.mk_addr (MemAddr.base l) e_zero
+let base_addr l = MemAddr.mk_addr (MemAddr.base l) F.e_zero
 let base_offset l = MemAddr.base_offset (MemAddr.base l) (MemAddr.offset l)
 let block_length sigma obj l =
-  let n_cells = e_div (allocated sigma l) (length_of_object obj) in
+  let n_cells = F.e_div (allocated sigma l) (length_of_object obj) in
   match obj with
   | C_comp ({ cfields = None } as c) ->
-    e_mul (Cvalues.bytes_length_of_opaque_comp c) n_cells
+    F.e_mul (Cvalues.bytes_length_of_opaque_comp c) n_cells
   | _ ->
-    e_fact (Ctypes.sizeof_object obj) n_cells
+    F.e_fact (Ctypes.sizeof_object obj) n_cells
 
 (* -------------------------------------------------------------------------- *)
 (* --- Cast                                                               --- *)
@@ -961,14 +965,16 @@ let int_of_loc _ l = MemAddr.int_of_addr l
 (* --- Frames                                                             --- *)
 (* -------------------------------------------------------------------------- *)
 
-let frames obj addr = function
-  | T_alloc -> []
-  | m ->
+let frames obj addr chunk =
+  match Sigma.ckind chunk with
+  | State.Mu T_alloc -> []
+  | State.Mu m ->
     let offset = length_of_object obj in
     let sizeof = F.e_one in
     let tau = Chunk.val_of_chunk m in
     let basename = Chunk.basename_of_chunk m in
     MemMemory.frames ~addr ~offset ~sizeof ~basename tau
+  | _ -> []
 
 (* -------------------------------------------------------------------------- *)
 (* --- Chunk element type                                                 --- *)
@@ -987,12 +993,12 @@ module ChunkContent = WpContext.Generator(Chunk)
       let generate c =
         let k = int_kind c in
         let p = Lang.freshvar ~basename:"m" (Chunk.tau_of_chunk c) in
-        let m = e_var p in
+        let m = F.e_var p in
         let name = Format.asprintf "is_%a_chunk" Ctypes.pp_int k in
         let lfun = Lang.generated_p ~coloring:true name in
         let l = Lang.freshvar ~basename:"l" (Lang.t_addr()) in
         let is_int = Cint.range k in
-        let def = p_forall [l] (is_int (F.e_get m (e_var l))) in
+        let def = F.p_forall [l] (is_int (F.e_get m (F.e_var l))) in
         Definitions.define_symbol {
           d_lfun = lfun ; d_types = 0 ;
           d_params = [p] ;
@@ -1004,15 +1010,15 @@ module ChunkContent = WpContext.Generator(Chunk)
       let compile = Lang.local generate
     end)
 
-let is_chunk sigma = function
-  | M_int _ as m ->
-    F.p_call (ChunkContent.get m) [ Sigma.value sigma m ]
-  | _ ->
-    p_true
-
 let is_well_formed sigma =
-  let open Sigma in
-  p_conj (Chunk.Set.fold (fun c l -> is_chunk sigma c :: l) (domain sigma) [])
+  let pool = ref [] in
+  Sigma.iter (fun c x ->
+      match Sigma.ckind c with
+      | State.Mu (M_int _ as chunk) ->
+        pool := F.p_call (ChunkContent.get chunk) [F.e_var x] :: !pool
+      | _ -> ()
+    ) sigma ;
+  F.p_conj !pool
 
 (* -------------------------------------------------------------------------- *)
 (* --- Loader                                                             --- *)
@@ -1020,10 +1026,8 @@ let is_well_formed sigma =
 
 module MODEL =
 struct
-  module Chunk = Chunk
-  module Sigma = Sigma
   let name = "MemTyped.LOADER"
-  type nonrec loc = loc
+  type loc = F.term
   let field = field
   let shift = shift
   let sizeof = length_of_object
@@ -1034,50 +1038,46 @@ struct
   let to_region_pointer l = 0,l
   let of_region_pointer _ _ l = l
 
-  let load_int sigma i l = F.e_get (Sigma.value sigma (M_int i)) l
-  let load_float sigma f l = F.e_get (Sigma.value sigma (m_float f)) l
-  let load_pointer sigma _t l = F.e_get (Sigma.value sigma M_pointer) l
+  let load_int sigma i l = F.e_get (State.value sigma (M_int i)) l
+  let load_float sigma f l = F.e_get (State.value sigma (m_float f)) l
+  let load_pointer sigma _t l = F.e_get (State.value sigma M_pointer) l
 
   let last sigma obj l =
     let n = length_of_object obj in
-    e_sub (F.e_div (allocated sigma l) n) e_one
+    F.e_sub (F.e_div (allocated sigma l) n) F.e_one
 
   let memcpy obj ~mtgt ~msrc ~ltgt ~lsrc ~length chunk =
-    if chunk <> T_alloc then
+    match Sigma.ckind chunk with
+    | State.Mu T_alloc -> msrc
+    | State.Mu _ ->
       let n = F.e_mul (length_of_object obj) length in
       F.e_fun f_memcpy [mtgt;msrc;ltgt;lsrc;n]
-    else msrc
+    | _ -> assert false
 
   let eqmem_forall obj loc _chunk m1 m2 =
     let xp = Lang.freshvar ~basename:"p" MemAddr.t_addr in
     let p = F.e_var xp in
     let n = length_of_object obj in
-    let separated = F.p_call MemAddr.p_separated [p;e_one;loc;n] in
-    let equal = p_equal (e_get m1 p) (e_get m2 p) in
+    let separated = F.p_call MemAddr.p_separated [p;F.e_one;loc;n] in
+    let equal = F.p_equal (F.e_get m1 p) (F.e_get m2 p) in
     [xp],separated,equal
 
-  let updated sigma c l v = c , F.e_set (Sigma.value sigma c) l v
+  let updated sigma c l v = State.chunk c , F.e_set (State.value sigma c) l v
 
   let store_int sigma i l v = updated sigma (M_int i) l v
   let store_float sigma f l v = updated sigma (m_float f) l v
   let store_pointer sigma _ty l v = updated sigma M_pointer l v
 
   let set_init_atom sigma _obj l v = updated sigma T_init l v
-  let is_init_atom sigma _ l = F.e_get (Sigma.value sigma T_init) l
+  let is_init_atom sigma _ l = F.e_get (State.value sigma T_init) l
 
   let is_init_range sigma obj loc length =
     let n = F.e_mul (length_of_object obj) length in
-    F.p_call p_is_init_r [ Sigma.value sigma T_init ; loc ; n ]
+    F.p_call p_is_init_r [ State.value sigma T_init ; loc ; n ]
 
   let set_init obj loc ~length _chunk ~current =
     let n = F.e_mul (length_of_object obj) length in
     F.e_fun f_set_init [current;loc;n]
-(*
-  let monotonic_init s1 s2 =
-    let m1 = Sigma.value s1 T_init in
-    let m2 = Sigma.value s2 T_init in
-    F.p_call p_monotonic [m1; m2]
-*)
 end
 
 module LOADER = MemLoader.Make(MODEL)
@@ -1093,7 +1093,7 @@ let domain = LOADER.domain
 
 let assigned seq obj loc =
   (* Maintain always initialized values initialized *)
-  Assert (MemMemory.cinits (Sigma.value seq.post T_init)) ::
+  Assert (MemMemory.cinits (State.value seq.post T_init)) ::
   LOADER.assigned seq obj loc
 
 (* -------------------------------------------------------------------------- *)
@@ -1105,15 +1105,15 @@ let loc_compare f_cmp i_cmp p q =
   | L.Yes -> i_cmp (MemAddr.offset p) (MemAddr.offset q)
   | L.Maybe | L.No -> f_cmp p q
 
-let is_null l = p_equal l null
-let loc_eq = p_equal
-let loc_neq = p_neq
-let loc_lt = loc_compare MemAddr.addr_lt p_lt
-let loc_leq = loc_compare MemAddr.addr_le p_leq
+let is_null l = F.p_equal l null
+let loc_eq = F.p_equal
+let loc_neq = F.p_neq
+let loc_lt = loc_compare MemAddr.addr_lt F.p_lt
+let loc_leq = loc_compare MemAddr.addr_le F.p_leq
 let loc_diff obj p q =
-  let delta = e_sub (MemAddr.offset p) (MemAddr.offset q) in
+  let delta = F.e_sub (MemAddr.offset p) (MemAddr.offset q) in
   let size = length_of_object obj in
-  e_div delta size
+  F.e_div delta size
 
 (* -------------------------------------------------------------------------- *)
 (* --- Validity                                                           --- *)
@@ -1125,17 +1125,17 @@ let s_valid sigma acs p n =
     | RD -> MemAddr.valid_rd
     | OBJ -> (fun m p _ -> MemAddr.valid_obj m p)
   in
-  valid (Sigma.value sigma T_alloc) p n
+  valid (State.value sigma T_alloc) p n
 
 let s_invalid sigma p n =
-  MemAddr.invalid (Sigma.value sigma T_alloc) p n
+  MemAddr.invalid (State.value sigma T_alloc) p n
 
 let segment phi = function
   | Rloc(obj,l) ->
     phi l (length_of_object obj)
   | Rrange(l,obj,Some a,Some b) ->
     let l = shift l obj a in
-    let n = e_mul (length_of_object obj) (e_range a b) in
+    let n = F.e_mul (length_of_object obj) (F.e_range a b) in
     phi l n
   | Rrange(l,_,a,b) ->
     Wp_parameters.abort ~current:true
@@ -1147,8 +1147,8 @@ let invalid sigma = segment (s_invalid sigma)
 
 let frame sigma =
   let wellformed_frame phi chunk =
-    if Sigma.mem sigma chunk
-    then [ phi (Sigma.value sigma chunk) ]
+    if State.mem sigma chunk
+    then [ phi (State.value sigma chunk) ]
     else []
   in
   wellformed_frame MemAddr.linked T_alloc @
@@ -1157,7 +1157,7 @@ let frame sigma =
   wellformed_frame MemMemory.framed M_pointer
 
 let alloc sigma xs =
-  if xs = [] then sigma else Sigma.havoc_chunk sigma T_alloc
+  if xs = [] then sigma else Sigma.havoc_chunk sigma (State.chunk T_alloc)
 
 let scope seq scope xs =
   if xs = [] then [] else
@@ -1165,13 +1165,13 @@ let scope seq scope xs =
       List.fold_left
         (fun m x ->
            let size = match scope with
-             | Sigs.Leave -> e_zero
-             | Sigs.Enter -> length_of_typ x.vtype
+             | Memory.Leave -> F.e_zero
+             | Memory.Enter -> length_of_typ x.vtype
            in F.e_set m (BASE.get x) size)
-        (Sigma.value seq.pre T_alloc) xs in
-    [ p_equal (Sigma.value seq.post T_alloc) alloc ]
+        (State.value seq.pre T_alloc) xs in
+    [ F.p_equal (State.value seq.post T_alloc) alloc ]
 
-let global _sigma p = p_leq MemAddr.(region @@ base p) e_zero
+let global _sigma p = F.p_leq MemAddr.(region @@ base p) F.e_zero
 
 (* -------------------------------------------------------------------------- *)
 (* --- Segments                                                           --- *)
@@ -1191,8 +1191,6 @@ let separated =
 (* --- State Model                                                        --- *)
 (* -------------------------------------------------------------------------- *)
 
-type state = chunk Tmap.t
-
 let rec lookup_a e =
   match F.repr e with
   | L.Fun( f , [e] ) when MemAddr.is_f_global f -> lookup_a e
@@ -1205,39 +1203,26 @@ and lookup_f f es =
     | RS_Index _ , [e;k] -> Mstate.index (lookup_lv e) k
     | _ -> raise Not_found
   with Not_found when es = [] ->
-    Sigs.(Mvar (RegisterBASE.find f),[])
+    Memory.(Mvar (RegisterBASE.find f),[])
 
-and lookup_lv e = try lookup_a e with Not_found -> Sigs.(Mmem e,[])
-
-let mchunk c =
-  match c with
-  | T_init -> Sigs.Mchunk (Pretty_utils.to_string Chunk.pretty c, KInit)
-  | _ -> Sigs.Mchunk (Pretty_utils.to_string Chunk.pretty c, KValue)
+and lookup_lv e = try lookup_a e with Not_found -> Memory.(Mmem e,[])
 
 let lookup s e =
-  try mchunk (Tmap.find e s)
-  with Not_found ->
-  try match F.repr e with
-    | L.Fun( f , es ) -> Sigs.Maddr (lookup_f f es)
-    | L.Aget( m , k ) when Tmap.find m s = T_init ->
-      Sigs.Mlval (lookup_lv k, KInit)
-    | L.Aget( m , k ) when Tmap.find m s <> T_alloc ->
-      Sigs.Mlval (lookup_lv k, KValue)
-    | _ -> Sigs.Mterm
-  with Not_found -> Sigs.Mterm
-
-let apply f s =
-  Tmap.fold (fun m c w -> Tmap.add (f m) c w) s Tmap.empty
-
-let iter f s = Tmap.iter (fun m c -> f (mchunk c) m) s
-
-let state (sigma : sigma) =
-  let s = ref Tmap.empty in
-  Sigma.iter (fun c x -> s := Tmap.add (e_var x) c !s) sigma ; !s
+  match F.repr e with
+  | L.Fun( f , es ) -> Memory.Maddr (lookup_f f es)
+  | L.Aget( m , k ) ->
+    begin
+      match Sigma.ckind @@ F.Tmap.find m s with
+      | State.Mu T_alloc -> raise Not_found
+      | State.Mu T_init -> Memory.Minit (lookup_lv k)
+      | State.Mu _ -> Memory.Mlval (lookup_lv k)
+      | _ -> raise Not_found
+    end
+  | _ -> raise Not_found
 
 let heap domain state =
-  Tmap.fold (fun m c w ->
-      if Vars.intersect (F.vars m) domain
+  F.Tmap.fold (fun m c w ->
+      if F.Vars.intersect (F.vars m) domain
       then Heap.Map.add c m w else w
     ) state Heap.Map.empty
 
@@ -1257,10 +1242,15 @@ let updates seq domain =
   let post = heap domain seq.post in
   Heap.Map.iter2
     (fun chunk v1 v2 ->
-       if chunk <> T_alloc then
-         match v1 , v2 with
-         | Some v1 , Some v2 -> pool := Bag.concat (diff v1 v2) !pool
-         | _ -> ())
+       match Sigma.ckind chunk with
+       | State.Mu mc ->
+         begin
+           if mc <> T_alloc then
+             match v1 , v2 with
+             | Some v1 , Some v2 -> pool := Bag.concat (diff v1 v2) !pool
+             | _ -> ()
+         end
+       | _ -> ())
     pre post ;
   !pool
 
diff --git a/src/plugins/wp/MemTyped.mli b/src/plugins/wp/MemTyped.mli
index f59c428b0dc8d07e0850ee66bf392e4aa92c70a9..fe26bfe0f9bb602d54f26d858978968ad17a1012 100644
--- a/src/plugins/wp/MemTyped.mli
+++ b/src/plugins/wp/MemTyped.mli
@@ -27,4 +27,4 @@
 type pointer = NoCast | Fits | Unsafe
 val pointer : pointer Context.value
 
-include Sigs.Model
+include Memory.Model
diff --git a/src/plugins/wp/MemVal.ml b/src/plugins/wp/MemVal.ml
index 8abf5fc6f52efcb1e80cb0316628d8b89495f1ca..cd725a6ae1344b1c2d806e73b923206ab8537086 100644
--- a/src/plugins/wp/MemVal.ml
+++ b/src/plugins/wp/MemVal.ml
@@ -28,8 +28,8 @@ open Cil_types
 open Cil_datatype
 open Ctypes
 open Lang
-open Lang.F
-open Sigs
+open Sigma
+open Memory
 open Definitions
 
 module Logic = Qed.Logic
@@ -63,13 +63,13 @@ sig
   val cvar : varinfo -> t
 
   val field : t -> Cil_types.fieldinfo -> t
-  val shift : t -> Ctypes.c_object -> term -> t
+  val shift : t -> Ctypes.c_object -> F.term -> t
   val base_addr : t -> t
 
   val load : state -> t -> Ctypes.c_object -> t
 
   val domain : t -> Base.t list
-  val offset : t -> (term -> pred)
+  val offset : t -> (F.term -> F.pred)
 
   val pretty : Format.formatter -> t -> unit
 end
@@ -100,17 +100,17 @@ let f_shift  = Lang.extern_f ~library ~result:t_addr "shift"
 let f_global = Lang.extern_f ~library ~result:t_addr "global"
 let f_null   = Lang.extern_f ~library ~result:t_addr "null"
 
-let a_null = F.constant (e_fun f_null []) (* { base = 0; offset = 0 } *)
-let a_base p = e_fun f_base [p]           (* p -> p.offset *)
-let a_offset p = e_fun f_offset [p]       (* p -> p.base *)
-let a_global b = e_fun f_global [b]       (* b -> { base = b; offset = 0 } *)
-let a_shift l k = e_fun f_shift [l;k]     (* p k -> { p w/ offset = p.offset + k } *)
+let a_null = F.constant (F.e_fun f_null []) (* { base = 0; offset = 0 } *)
+let a_base p = F.e_fun f_base [p]           (* p -> p.offset *)
+let a_offset p = F.e_fun f_offset [p]       (* p -> p.base *)
+let a_global b = F.e_fun f_global [b]       (* b -> { base = b; offset = 0 } *)
+let a_shift l k = F.e_fun f_shift [l;k]     (* p k -> { p w/ offset = p.offset + k } *)
 let a_addr b k = a_shift (a_global b) k   (* b k -> { base = b; offset = k } *)
 
 (* -------------------------------------------------------------------------- *)
 (* ---  Cmath Wrapper                                                     --- *)
 (* -------------------------------------------------------------------------- *)
-let a_iabs i = e_fun ~result:Logic.Int Cmath.f_iabs [i]    (* x -> |x| *)
+let a_iabs i = F.e_fun ~result:Logic.Int Cmath.f_iabs [i]    (* x -> |x| *)
 
 (* -------------------------------------------------------------------------- *)
 (* ---  MemValue Types                                                     --- *)
@@ -127,7 +127,7 @@ let phi_base t = match F.repr t with
   | _ -> raise Not_found
 
 let phi_offset t = match F.repr t with
-  | Logic.Fun (f, [p; k]) when f == f_shift -> e_add (a_offset p) k
+  | Logic.Fun (f, [p; k]) when f == f_shift -> F.e_add (a_offset p) k
   | Logic.Fun (f, _) when f == f_global -> F.e_zero
   | _ -> raise Not_found
 
@@ -187,8 +187,7 @@ struct
     | Unknown (_, _, m) -> Integer.succ m
     | Variable { max_allocable } -> Integer.succ max_allocable
 
-  let size_from_validity b =
-    Integer.(e_div (bitsize_from_validity b) eight)
+  let size_from_validity b = Integer.(e_div (bitsize_from_validity b) eight)
 end
 
 
@@ -225,68 +224,57 @@ struct
   (* -------------------------------------------------------------------------- *)
   (* ---  Chunk                                                             --- *)
   (* -------------------------------------------------------------------------- *)
-  type chunk =
-    | M_base of Base.t
 
   module Chunk =
   struct
-    type t = chunk
+    type t = Base.t
     let self = "MemVal.Chunk"
-    let hash = function
-      | M_base b -> 5 * Base.hash b
-    let equal c1 c2 = match c1, c2 with
-      | M_base b1, M_base b2 -> Base.equal b1 b2
-    let compare c1 c2 = match c1, c2 with
-      | M_base b1, M_base b2 -> Base.compare b1 b2
-    let pretty fmt = function
-      | M_base b -> Base.pretty fmt b
-    let tau_of_chunk = function
-      | M_base _ -> t_words
-    let basename_of_base = function
+    let hash = Base.hash
+    let equal = Base.equal
+    let compare = Base.compare
+    let pretty = Base.pretty
+    let tau_of_chunk _base = t_words
+    let basename_of_chunk = function
       | Base.Var (vi, _) -> Format.sprintf "MVar_%s" (LogicUsage.basename vi)
       | Base.CLogic_Var (_, _, _) -> assert false (* not supposed to append. *)
       | Base.Null -> "MNull"
       | Base.String (eid, _) -> Format.sprintf "MStr_%d" eid
       | Base.Allocated (vi, _dealloc, _) ->
         Format.sprintf "MAlloc_%s" (LogicUsage.basename vi)
-    let basename_of_chunk = function
-      | M_base b -> basename_of_base b
-    let is_framed = function
-      | M_base b ->
-        try
-          (match WpContext.get_scope () with
-           | WpContext.Global -> assert false
-           | WpContext.Kf kf -> Base.is_formal_or_local b (Kernel_function.get_definition kf))
-        with Invalid_argument _ | Kernel_function.No_Definition ->
-          assert false (* by context. *)
+
+    let is_init _ = false
+    let is_primary = function Base.Var _ -> true | _ -> false
+
+    let is_framed b =
+      try
+        (match WpContext.get_scope () with
+         | WpContext.Global -> assert false
+         | WpContext.Kf kf -> Base.is_formal_or_local b (Kernel_function.get_definition kf))
+      with Invalid_argument _ | Kernel_function.No_Definition ->
+        assert false (* by context. *)
+
   end
 
   let cluster () = Definitions.cluster ~id:"MemVal"  ()
 
-  module Heap = Qed.Collection.Make(Chunk)
-  module Sigma = Sigma.Make(Chunk)(Heap)
+  module State = Sigma.Make(Chunk)
 
   type loc = {
     loc_v : V.t;
-    loc_t : term (* of type addr *)
+    loc_t : F.term (* of type addr *)
   }
 
-  type sigma = Sigma.t
   type segment = loc rloc
 
-  type state = unit
-  let state _ = ()
-  let iter _ _ = ()
-  let lookup _ _ = Mterm
+  let lookup _ _ = raise Not_found
   let updates _ _ = Bag.empty
-  let apply _ _ = ()
 
   let pretty fmt l =
     Format.fprintf fmt "([@ t:%a,@ v:%a @])"
       F.pp_term l.loc_t
       V.pretty l.loc_v
 
-  let vars _l = Vars.empty
+  let vars _l = F.Vars.empty
   let occurs _x _l = false
 
   (* -------------------------------------------------------------------------- *)
@@ -365,14 +353,14 @@ struct
         let axiomatize ~obj:_ suffix t_mem t_data f_rd f_wr =
           let name = "axiom_" ^ suffix in
           let xw = Lang.freshvar ~basename:"w" t_mem in
-          let w = e_var xw in
+          let w = F.e_var xw in
           let xo = Lang.freshvar ~basename:"o" Logic.Int in
-          let o = e_var xo in
+          let o = F.e_var xo in
           let xv = Lang.freshvar ~basename:"v" t_data in
-          let v = e_var xv in
-          let p_write = e_fun f_wr [w; o; v] ~result:t_mem in
-          let p_read = e_fun f_rd [p_write; o] ~result:t_data in
-          let lemma = p_equal p_read v in
+          let v = F.e_var xv in
+          let p_write = F.e_fun f_wr [w; o; v] ~result:t_mem in
+          let p_read = F.e_fun f_rd [p_write; o] ~result:t_data in
+          let lemma = F.p_equal p_read v in
           let cluster = cluster () in
           (* if not (Wp_parameters.debug_atleast 1) then begin
            *   F.set_builtin_2 f_rd (phi_read ~obj ~read:f_rd ~write:f_wr)
@@ -389,21 +377,21 @@ struct
         let axiomatize2 ~obj suffix t_mem t_data f_rd f_wr =
           let name = "axiom_" ^ suffix ^ "_2" in
           let xw = Lang.freshvar ~basename:"w" t_mem in
-          let w = e_var xw in
+          let w = F.e_var xw in
           let xwo = Lang.freshvar ~basename:"xwo" Logic.Int in
-          let wo = e_var xwo in
+          let wo = F.e_var xwo in
           let xro = Lang.freshvar ~basename:"xro" Logic.Int in
-          let ro = e_var xro in
+          let ro = F.e_var xro in
           let xv = Lang.freshvar ~basename:"v" t_data in
-          let v = e_var xv in
-          let p_write = e_fun f_wr [w; wo; v] ~result:t_mem in
-          let p_read = e_fun f_rd [p_write; ro] ~result:t_data in
+          let v = F.e_var xv in
+          let p_write = F.e_fun f_wr [w; wo; v] ~result:t_mem in
+          let p_read = F.e_fun f_rd [p_write; ro] ~result:t_data in
           let sizeof = (F.e_int (Ctypes.sizeof_object obj)) in
           let offset = a_iabs (F.e_sub ro wo) in
           let lemma =
             F.p_imply
               (F.p_leq sizeof offset)
-              (F.p_equal p_read (e_fun f_rd [w; ro] ~result:t_data))
+              (F.p_equal p_read (F.e_fun f_rd [w; ro] ~result:t_data))
           in
           let cluster = cluster () in
           Definitions.define_lemma {
@@ -513,15 +501,15 @@ struct
     let d = V.domain l.loc_v in
     assert (d <> []);
     List.fold_left
-      (fun acc b -> Heap.Set.add (M_base b) acc)
-      Heap.Set.empty d
+      (fun acc b -> Domain.add (State.chunk b) acc)
+      Domain.empty d
 
   (* -------------------------------------------------------------------------- *)
   (* ---  Memory Load                                                       --- *)
   (* -------------------------------------------------------------------------- *)
   let load_value sigma obj l =
     let load_base base =
-      let mem = Sigma.value sigma (M_base base) in
+      let mem = State.value sigma base in
       let offset = a_offset l.loc_t in
       read obj ~mem ~offset
     in
@@ -534,7 +522,7 @@ struct
 
   let load_loc ~assume sigma obj l =
     let load_base v' base =
-      let mem = Sigma.value sigma (M_base base) in
+      let mem = State.value sigma base in
       let offset = a_offset l.loc_t in
       let rd = read obj ~mem ~offset in
       if assume then begin
@@ -550,7 +538,7 @@ struct
       loc_t = t;
     }
 
-  let load : sigma -> c_object -> loc -> loc value = fun sigma obj l ->
+  let load sigma obj l =
     StateRef.update ();
     begin match obj with
       | C_int _ | C_float _ -> load_value sigma obj l
@@ -559,15 +547,15 @@ struct
     end
 
   let load_init _sigma obj _loc =
-    e_var @@ Lang.freshvar ~basename:"i" @@ Lang.init_of_object obj
+    F.e_var @@ Lang.freshvar ~basename:"i" @@ Lang.init_of_object obj
 
   (* -------------------------------------------------------------------------- *)
   (* ---  Memory Store                                                      --- *)
   (* -------------------------------------------------------------------------- *)
-  let stored : sigma sequence -> c_object -> loc -> term -> equation list = fun seq obj l v ->
+  let stored seq obj l v =
     let mk_write cond base =
-      let wpre = Sigma.value seq.pre (M_base base) in
-      let wpost = Sigma.value seq.post (M_base base) in
+      let wpre = State.value seq.pre base in
+      let wpost = State.value seq.post base in
       let write = write obj ~mem:wpre ~offset:(a_offset l.loc_t) ~value:v in
       F.p_equal wpost (F.e_if cond write wpre)
     in
@@ -588,8 +576,8 @@ struct
 
   let copied seq obj ll lr =
     let v = match load seq.pre obj lr with
-      | Sigs.Val v -> v
-      | Sigs.Loc l -> l.loc_t
+      | Memory.Val v -> v
+      | Memory.Loc l -> l.loc_t
     in
     stored seq obj ll v
 
@@ -601,7 +589,7 @@ struct
   (* -------------------------------------------------------------------------- *)
   (* ---  Pointer Comparison                                                --- *)
   (* -------------------------------------------------------------------------- *)
-  let is_null l = p_equal l.loc_t a_null
+  let is_null l = F.p_equal l.loc_t a_null
 
   let loc_delta l1 l2 =
     match F.is_equal (a_base l1.loc_t) (a_base l2.loc_t) with
@@ -623,7 +611,7 @@ struct
   (* -------------------------------------------------------------------------- *)
 
   type range =
-    | LOC of loc * term (*size*)
+    | LOC of loc * F.term (*size*)
     | RANGE of loc * Vset.set (* offset range access from *loc* *)
 
   let range_of_rloc = function
@@ -631,7 +619,7 @@ struct
       LOC (l, F.e_int (Ctypes.sizeof_object obj))
     | Rrange (l, obj, Some a, Some b) ->
       let la = shift l obj a in
-      let n = e_fact (Ctypes.sizeof_object obj) (F.e_range a b) in
+      let n = F.e_fact (Ctypes.sizeof_object obj) (F.e_range a b) in
       LOC (la, n)
     | Rrange (l, obj, a_opt, b_opt) ->
       let f = F.e_fact (Ctypes.sizeof_object obj) in
@@ -657,7 +645,7 @@ struct
       Vset.range (Some F.e_zero) (Some mn_valid)
     | Base.Unknown (_, None, _) -> Vset.empty
 
-  let valid_range : sigma -> acs -> range -> pred = fun _ acs r ->
+  let valid_range acs r =
     let for_writing = match acs with RW -> true | RD -> false
                                    | OBJ -> true (* TODO:  *) in
     let l, base_offset = match r with
@@ -679,17 +667,16 @@ struct
   (** [valid sigma acs seg] returns the formula that tests if a given memory
       segment [seg] (in bytes) is valid (according to [acs]) at memory state
       [sigma]. **)
-  let valid : sigma -> acs -> segment -> pred = fun s acs seg ->
-    valid_range s acs (range_of_rloc seg)
+  let valid _sigma acs seg = valid_range acs (range_of_rloc seg)
 
   let invalid = fun _ _ -> F.p_true (* TODO *)
 
   (* -------------------------------------------------------------------------- *)
   (* ---  Scope                                                             --- *)
   (* -------------------------------------------------------------------------- *)
-  let alloc_sigma : sigma -> varinfo list -> sigma = fun sigma xs ->
+  let alloc_vars sigma xs =
     let alloc sigma x =
-      let havoc s c = Sigma.havoc_chunk s (M_base c) in
+      let havoc s b = Sigma.havoc_chunk s (State.chunk b) in
       let v = V.cvar x in
       List.fold_left havoc sigma (V.domain v)
     in
@@ -698,9 +685,9 @@ struct
   let alloc_pred _ _ _ = []
 
   let alloc sigma xs =
-    if xs = [] then sigma else alloc_sigma sigma xs
+    if xs = [] then sigma else alloc_vars sigma xs
 
-  let scope : sigma sequence -> scope -> varinfo list -> pred list = fun seq scope xs ->
+  let scope seq scope xs =
     match scope with
     | Enter -> []
     | Leave ->
@@ -712,13 +699,12 @@ struct
       Sigma.pretty seq.pre
       Sigma.pretty seq.post
       (Pretty_utils.pp_iter ~sep:" " List.iter Varinfo.pretty) xs
-      (Pretty_utils.pp_iter ~sep:" " List.iter pp_pred) preds;
+      (Pretty_utils.pp_iter ~sep:" " List.iter F.pp_pred) preds;
     preds
 
-  let global : sigma -> term (*addr*) -> pred = fun _ _ ->
+  let global _ _ =
     F.p_true (* True is harmless and WP never call this function... *)
 
-
   (* -------------------------------------------------------------------------- *)
   (* ---  Separation                                                        --- *)
   (* -------------------------------------------------------------------------- *)
@@ -729,26 +715,25 @@ struct
       l, Vset.range (Some a) (Some b)
     | RANGE (l, r) -> l, Vset.lift_add (Vset.singleton l.loc_t) r
 
-  let included : segment -> segment -> pred = fun s1 s2 ->
+  let included : segment -> segment -> F.pred = fun s1 s2 ->
     (* (b1 = b2) -> (r1 \in r2) *)
     let l1, vs1 = range_to_base_offset (range_of_rloc s1) in
     let l2, vs2 = range_to_base_offset (range_of_rloc s2) in
-    p_and
-      (p_equal (a_base l1.loc_t) (a_base l2.loc_t))
+    F.p_and
+      (F.p_equal (a_base l1.loc_t) (a_base l2.loc_t))
       (Vset.subset vs1 vs2)
 
-  let separated : segment -> segment -> pred = fun s1 s2 ->
+  let separated : segment -> segment -> F.pred = fun s1 s2 ->
     (* (b1 = b2) -> (r1 \cap r2) = \empty *)
     let l1, vs1 = range_to_base_offset (range_of_rloc s1) in
     let l2, vs2 = range_to_base_offset (range_of_rloc s2) in
-    p_and
-      (p_equal (a_base l1.loc_t) (a_base l2.loc_t))
+    F.p_and
+      (F.p_equal (a_base l1.loc_t) (a_base l2.loc_t))
       (Vset.disjoint vs1 vs2)
 
   let initialized _sigma _l = F.p_true (* todo *)
   let is_well_formed _ = F.p_true (* todo *)
   let base_offset _loc = assert false (* TODO *)
-  type domain = Sigma.domain
   let no_binder = { bind = fun _ f v -> f v }
   let configure_ia _ = no_binder (* todo *)
   let hypotheses x = x (* todo *)
@@ -756,9 +741,6 @@ struct
 
 end
 
-
-
-
 (* -------------------------------------------------------------------------- *)
 (* ---  EVA Instance                                                      --- *)
 (* -------------------------------------------------------------------------- *)
diff --git a/src/plugins/wp/MemVal.mli b/src/plugins/wp/MemVal.mli
index 1526f08395dfa76e90a5da63871489d3fea70fa9..f03da645528751f6973507643199d4fed296d0c4 100644
--- a/src/plugins/wp/MemVal.mli
+++ b/src/plugins/wp/MemVal.mli
@@ -88,7 +88,7 @@ sig
   val pretty : Format.formatter -> t -> unit
 end
 
-module Make(_ : Value) : Sigs.Model
+module Make(_ : Value) : Memory.Model
 
 (** The glue between WP and EVA. **)
 module Eva : Value
diff --git a/src/plugins/wp/MemVar.ml b/src/plugins/wp/MemVar.ml
index afbf8f5ab9bb3db9d52e46376e966b8fcd6ef601..ac5e5a73fd0e674fba9d2a298ba6bfe8810d1774 100644
--- a/src/plugins/wp/MemVar.ml
+++ b/src/plugins/wp/MemVar.ml
@@ -31,7 +31,7 @@ open Ctypes
 open MemoryContext
 open Lang
 open Lang.F
-open Sigs
+open Memory
 
 module type VarUsage =
 sig
@@ -64,7 +64,7 @@ struct
   let iter = Raw.iter
 end
 
-module Make(V : VarUsage)(M : Sigs.Model) =
+module Make(V : VarUsage)(M : Memory.Model) =
 struct
 
   (* -------------------------------------------------------------------------- *)
@@ -89,12 +89,6 @@ struct
   (* ---  Chunk                                                             --- *)
   (* -------------------------------------------------------------------------- *)
 
-  type chunk =
-    | Var of varinfo
-    | Alloc of varinfo
-    | Init of varinfo
-    | Mem of M.Chunk.t
-
   let is_framed_var x = not x.vglob && not x.vaddrof
   (* Can not use VarUsage info, since (&x) can still be passed
      to the function and be modified by the call (when it assigns everything). *)
@@ -113,6 +107,8 @@ struct
       | _ -> x.vtype
     let tau_of_chunk x = Lang.tau_of_ctype (typ_of_chunk x)
     let is_framed = is_framed_var
+    let is_init _ = false
+    let is_primary _ = true
     let basename_of_chunk = LogicUsage.basename
   end
 
@@ -131,6 +127,8 @@ struct
         "ra_" ^ LogicUsage.basename x
       | NotUsed | ByValue | ByShift | ByAddr | InContext _ | InArray _ ->
         "ta_" ^ LogicUsage.basename x
+    let is_init _ = false
+    let is_primary _ = false
     let is_framed = is_framed_var
   end
 
@@ -147,262 +145,19 @@ struct
       | ByRef -> Cil.typeOf_pointed x.vtype
       | _ -> x.vtype
     let tau_of_chunk x = Lang.init_of_ctype (typ_of_chunk x)
+    let is_init _ = true
+    let is_primary _ = false
     let is_framed = is_framed_var
     let basename_of_chunk x = "Init_" ^ (LogicUsage.basename x)
   end
 
-  module Chunk =
-  struct
-    type t = chunk
-    let self = "varmem"
-    let hash = function
-      | Var x -> 3 * Varinfo.hash x
-      | Alloc x -> 5 * Varinfo.hash x
-      | Mem m -> 7 * M.Chunk.hash m
-      | Init x -> 11 * Varinfo.hash x
-    let compare c1 c2 =
-      if c1 == c2 then 0 else
-        match c1 , c2 with
-        | Var x , Var y
-        | Alloc x , Alloc y
-        | Init x, Init y -> Varinfo.compare x y
-        | Mem p , Mem q -> M.Chunk.compare p q
-        | Var _ , _ -> (-1)
-        | _ , Var _ -> 1
-        | Init _, _ -> (-1)
-        | _, Init _ -> 1
-        | Alloc _  , _ -> (-1)
-        | _ , Alloc _ -> 1
-    let equal c1 c2 = (compare c1 c2 = 0)
-    let pretty fmt = function
-      | Var x -> Varinfo.pretty fmt x
-      | Alloc x -> Format.fprintf fmt "alloc(%a)" Varinfo.pretty x
-      | Init x -> Format.fprintf fmt "init(%a)" Varinfo.pretty x
-      | Mem m -> M.Chunk.pretty fmt m
-    let tau_of_chunk = function
-      | Var x -> VAR.tau_of_chunk x
-      | Alloc x -> VALLOC.tau_of_chunk x
-      | Init x -> VINIT.tau_of_chunk x
-      | Mem m -> M.Chunk.tau_of_chunk m
-    let basename_of_chunk = function
-      | Var x -> VAR.basename_of_chunk x
-      | Alloc x -> VALLOC.basename_of_chunk x
-      | Init x -> VINIT.basename_of_chunk x
-      | Mem m -> M.Chunk.basename_of_chunk m
-    let is_framed = function
-      | Var x -> VAR.is_framed x
-      | Alloc x -> VALLOC.is_framed x
-      | Init x -> VINIT.is_framed x
-      | Mem m -> M.Chunk.is_framed m
-  end
-
   (* -------------------------------------------------------------------------- *)
   (* ---  Sigma                                                             --- *)
   (* -------------------------------------------------------------------------- *)
 
-  module HEAP = Qed.Collection.Make(VAR)
-  module TALLOC = Qed.Collection.Make(VALLOC)
-  module TINIT = Qed.Collection.Make(VINIT)
-  module SIGMA = Sigma.Make(VAR)(HEAP)
-  module ALLOC = Sigma.Make(VALLOC)(TALLOC)
-  module INIT = Sigma.Make(VINIT)(TINIT)
-  module Heap = Qed.Collection.Make(Chunk)
-
-  type sigma = {
-    mem : M.Sigma.t ;
-    vars : SIGMA.t ;
-    init : INIT.t ;
-    alloc : ALLOC.t ;
-  }
-
-  module Sigma =
-  struct
-    type t = sigma
-    type chunk = Chunk.t
-    module Chunk = Heap
-    type domain = Heap.set
-    let empty = Heap.Set.empty
-    let union = Heap.Set.union
-
-    let create () = {
-      vars = SIGMA.create () ;
-      init = INIT.create () ;
-      alloc = ALLOC.create () ;
-      mem = M.Sigma.create () ;
-    }
-
-    let copy s = {
-      vars = SIGMA.copy s.vars ;
-      init = INIT.copy s.init ;
-      alloc = ALLOC.copy s.alloc ;
-      mem = M.Sigma.copy s.mem ;
-    }
-
-    let choose s1 s2 =
-      let s = SIGMA.choose s1.vars s2.vars in
-      let i = INIT.choose s1.init s2.init in
-      let a = ALLOC.choose s1.alloc s2.alloc in
-      let m = M.Sigma.choose s1.mem s2.mem in
-      { vars = s ; alloc = a ; mem = m ; init = i }
-
-    let merge s1 s2 =
-      let s,pa1,pa2 = SIGMA.merge s1.vars s2.vars in
-      let i,ia1,ia2 = INIT.merge s1.init s2.init in
-      let a,ta1,ta2 = ALLOC.merge s1.alloc s2.alloc in
-      let m,qa1,qa2 = M.Sigma.merge s1.mem s2.mem in
-      { vars = s ; alloc = a ; mem = m ; init = i } ,
-      Passive.union (Passive.union (Passive.union pa1 ta1) qa1) ia1 ,
-      Passive.union (Passive.union (Passive.union pa2 ta2) qa2) ia2
-
-    let merge_list l =
-      let s,pa = SIGMA.merge_list (List.map (fun s -> s.vars) l) in
-      let i,ia = INIT.merge_list (List.map (fun s -> s.init) l) in
-      let a,ta = ALLOC.merge_list (List.map (fun s -> s.alloc) l) in
-      let m,qa = M.Sigma.merge_list (List.map (fun s -> s.mem) l) in
-      { vars = s ; alloc = a ; mem = m ; init = i } ,
-      let union = List.map2 Passive.union in
-      union (union (union pa ta) qa) ia
-
-    let join s1 s2 =
-      Passive.union
-        (Passive.union
-           (Passive.union
-              (SIGMA.join s1.vars s2.vars)
-              (ALLOC.join s1.alloc s2.alloc))
-           (M.Sigma.join s1.mem s2.mem))
-        (INIT.join s1.init s2.init)
-
-    let get s = function
-      | Var x -> SIGMA.get s.vars x
-      | Init x -> INIT.get s.init x
-      | Alloc x -> ALLOC.get s.alloc x
-      | Mem m -> M.Sigma.get s.mem m
-
-    let mem s = function
-      | Var x -> SIGMA.mem s.vars x
-      | Init x -> INIT.mem s.init x
-      | Alloc x -> ALLOC.mem s.alloc x
-      | Mem m -> M.Sigma.mem s.mem m
-
-    let value s c = e_var (get s c)
-
-    let iter f s =
-      begin
-        SIGMA.iter (fun x -> f (Var x)) s.vars ;
-        INIT.iter (fun x -> f (Init x)) s.init ;
-        ALLOC.iter (fun x -> f (Alloc x)) s.alloc ;
-        M.Sigma.iter (fun m -> f (Mem m)) s.mem ;
-      end
-
-    let iter2 f s t =
-      begin
-        SIGMA.iter2 (fun x a b -> f (Var x) a b) s.vars t.vars ;
-        INIT.iter2 (fun x a b -> f (Init x) a b) s.init t.init ;
-        ALLOC.iter2 (fun x a b -> f (Alloc x) a b) s.alloc t.alloc ;
-        M.Sigma.iter2 (fun m p q -> f (Mem m) p q) s.mem t.mem ;
-      end
-
-    let domain_partition r =
-      begin
-        let xs = ref HEAP.Set.empty in
-        let is = ref TINIT.Set.empty in
-        let ts = ref TALLOC.Set.empty in
-        let ms = ref M.Heap.Set.empty in
-        Heap.Set.iter
-          (function
-            | Var x -> xs := HEAP.Set.add x !xs
-            | Init x -> is := TINIT.Set.add x !is
-            | Alloc x -> ts := TALLOC.Set.add x !ts
-            | Mem c -> ms := M.Heap.Set.add c !ms
-          ) r ;
-        !xs , !is, !ts , !ms
-      end
-
-    let domain_var xs =
-      HEAP.Set.fold (fun x s -> Heap.Set.add (Var x) s) xs Heap.Set.empty
-
-    let domain_init xs =
-      TINIT.Set.fold (fun x s -> Heap.Set.add (Init x) s) xs Heap.Set.empty
-
-    let domain_alloc ts =
-      TALLOC.Set.fold (fun x s -> Heap.Set.add (Alloc x) s) ts Heap.Set.empty
-
-    let domain_mem ms =
-      M.Heap.Set.fold (fun m s -> Heap.Set.add (Mem m) s) ms Heap.Set.empty
-
-    let assigned ~pre ~post w =
-      let w_vars , w_init, w_alloc , w_mem = domain_partition w in
-      let h_vars = SIGMA.assigned ~pre:pre.vars ~post:post.vars w_vars in
-      let h_init = INIT.assigned ~pre:pre.init ~post:post.init w_init in
-      let h_alloc = ALLOC.assigned ~pre:pre.alloc ~post:post.alloc w_alloc in
-      let h_mem = M.Sigma.assigned ~pre:pre.mem ~post:post.mem w_mem in
-      Bag.ulist [h_vars;h_init;h_alloc;h_mem]
-
-    let havoc s r =
-      let rvar , rinit, ralloc , rmem = domain_partition r
-      in {
-        vars = SIGMA.havoc s.vars rvar ;
-        init = INIT.havoc s.init rinit ;
-        alloc = ALLOC.havoc s.alloc ralloc ;
-        mem = M.Sigma.havoc s.mem rmem ;
-      }
-
-    let havoc_chunk s = function
-      | Var x -> { s with vars = SIGMA.havoc_chunk s.vars x }
-      | Init x -> { s with init = INIT.havoc_chunk s.init x }
-      | Alloc x -> { s with alloc = ALLOC.havoc_chunk s.alloc x }
-      | Mem m -> { s with mem = M.Sigma.havoc_chunk s.mem m }
-
-    let havoc_any ~call s = {
-      alloc = s.alloc ;
-      vars = SIGMA.havoc_any ~call s.vars ;
-      init = INIT.havoc_any ~call s.init ;
-      mem = M.Sigma.havoc_any ~call s.mem ;
-    }
-
-    let remove_chunks s r =
-      let rvar , rinit, ralloc , rmem = domain_partition r
-      in {
-        vars = SIGMA.remove_chunks s.vars rvar ;
-        init = INIT.remove_chunks s.init rinit ;
-        alloc = ALLOC.remove_chunks s.alloc ralloc ;
-        mem = M.Sigma.remove_chunks s.mem rmem ;
-      }
-
-    let domain s =
-      Heap.Set.union
-        (Heap.Set.union
-           (Heap.Set.union
-              (domain_var (SIGMA.domain s.vars))
-              (domain_alloc (ALLOC.domain s.alloc)))
-           (domain_mem (M.Sigma.domain s.mem)))
-        (domain_init (INIT.domain s.init))
-
-    let writes s =
-      Heap.Set.union
-        (Heap.Set.union
-           (Heap.Set.union
-              (domain_var (SIGMA.writes {pre=s.pre.vars;post=s.post.vars}))
-              (domain_alloc (ALLOC.writes {pre=s.pre.alloc;post=s.post.alloc})))
-           (domain_mem (M.Sigma.writes {pre=s.pre.mem;post=s.post.mem})))
-        (domain_init (INIT.writes {pre=s.pre.init;post=s.post.init}))
-
-    let pretty fmt s =
-      Format.fprintf fmt "@[<hov 2>{X:@[%a@]@ I:@[%a@]@ T:@[%a@]@ M:@[%a@]}@]"
-        SIGMA.pretty s.vars
-        INIT.pretty s.init
-        ALLOC.pretty s.alloc
-        M.Sigma.pretty s.mem
-
-  end
-
-  type domain = Sigma.domain
-
-  let get_var s x = SIGMA.get s.vars x
-  let get_term s x = e_var (get_var s x)
-
-  let get_init s x = INIT.get s.init x
-  let get_init_term s x = e_var (get_init s x)
+  module SVAR = Sigma.Make(VAR)
+  module SALLOC = Sigma.Make(VALLOC)
+  module SINIT = Sigma.Make(VINIT)
 
   (* -------------------------------------------------------------------------- *)
   (* ---  State Pretty Printer                                              --- *)
@@ -410,11 +165,6 @@ struct
 
   type ichunk = Iref of varinfo | Ivar of varinfo | Iinit of varinfo
 
-  type state = {
-    svar : ichunk Tmap.t ;
-    smem : M.state ;
-  }
-
   module IChunk =
   struct
 
@@ -449,55 +199,41 @@ struct
 
   module Icmap = Qed.Mergemap.Make(IChunk)
 
-  let set_chunk v c m =
-    let c =
-      try
-        let c0 = Tmap.find v m in
-        if IChunk.compare c c0 < 0 then c else c0
-      with Not_found -> c in
-    Tmap.add v c m
-
-  let state s =
-    let m = ref Tmap.empty in
-    SIGMA.iter (fun x v ->
-        let c = match V.param x with ByRef -> Iref x | _ -> Ivar x in
-        m := set_chunk (e_var v) c !m
-      ) s.vars ;
-    INIT.iter (fun x v ->
-        match V.param x with
-        | ByRef -> ()
-        | _ -> m := set_chunk (e_var v) (Iinit x) !m
-      ) s.init ;
-    { svar = !m ; smem = M.state s.mem }
-
   let ilval = function
     | Iref x -> (Mvar x,[Mindex e_zero])
     | Ivar x | Iinit x -> (Mvar x,[])
 
   let imval c = match c with
-    | Iref _ | Ivar _ -> Sigs.Mlval (ilval c, KValue)
-    | Iinit _ -> Sigs.Mlval (ilval c, KInit)
+    | Iref _ | Ivar _ -> Memory.Mlval (ilval c)
+    | Iinit _ -> Memory.Minit (ilval c)
 
-  let lookup s e =
-    try imval (Tmap.find e s.svar)
-    with Not_found -> M.lookup s.smem e
+  let is_ref x =
+    match V.param x with
+    | InContext _ | InArray _ | ByRef -> true
+    | _ -> false
 
-  let apply f s =
-    let m = ref Tmap.empty in
-    Tmap.iter (fun e c ->
-        let e = f e in
-        m := set_chunk e c !m ;
-      ) s.svar ;
-    { svar = !m ; smem = M.apply f s.smem }
+  let iref x = if is_ref x then Iref x else Ivar x
 
-  let iter f s =
-    Tmap.iter (fun v c -> f (imval c) v) s.svar ;
-    M.iter f s.smem
+  let lookup (s : Sigma.state) e =
+    try
+      match Sigma.ckind @@ Tmap.find e s with
+      | SVAR.Mu x -> imval (iref x)
+      | SINIT.Mu x -> imval (Iinit x)
+      | SALLOC.Mu _ -> raise Not_found
+      | _ -> raise Not_found
+    with Not_found ->
+      M.lookup s e
 
   let icmap domain istate =
-    Tmap.fold (fun m c w ->
-        if Vars.intersect (F.vars m) domain
-        then Icmap.add c m w else w
+    Tmap.fold
+      (fun m c w ->
+         if Vars.intersect (F.vars m) domain
+         then
+           match Sigma.ckind c with
+           | SVAR.Mu x -> Icmap.add (iref x) m w
+           | SINIT.Mu x -> Icmap.add (Iinit x) m w
+           | _ -> w
+         else w
       ) istate Icmap.empty
 
   let rec diff lv v1 v2 =
@@ -521,9 +257,9 @@ struct
     | (Lang.Mfield _,_)::_ -> Bag.elt (Mstore(lv,v2))
     | [] -> Bag.empty
 
-  let updates seq domain =
-    let pre = icmap domain seq.pre.svar in
-    let post = icmap domain seq.post.svar in
+  let updates (seq : Sigma.state sequence) domain =
+    let pre = icmap domain seq.pre in
+    let post = icmap domain seq.post in
     let pool = ref Bag.empty in
     Icmap.iter2
       (fun c v1 v2 ->
@@ -532,8 +268,7 @@ struct
          | None , Some v -> pool := Bag.add (Mstore(ilval c,v)) !pool
          | Some v1 , Some v2 -> pool := Bag.concat (diff (ilval c) v1 v2) !pool
       ) pre post ;
-    let seq_mem =  { pre = seq.pre.smem ; post = seq.post.smem } in
-    Bag.concat !pool (M.updates seq_mem domain)
+    Bag.concat !pool (M.updates seq domain)
 
   (* -------------------------------------------------------------------------- *)
   (* ---  Location                                                          --- *)
@@ -699,8 +434,6 @@ struct
     | Field f -> M.field l f
     | Shift(e,k) -> M.shift l e k
 
-  let mseq_of_seq seq = { pre = seq.pre.mem ; post = seq.post.mem }
-
   let mloc_of_path m x ofs =
     let l = match m with
       | CTXT Nullable | CARR Nullable -> M.pointer_loc @@ nullable_address x
@@ -745,7 +478,7 @@ struct
     | Val(_,_,ofs) -> List.fold_left byte_offset e_zero ofs
 
   let block_length sigma obj = function
-    | Loc l -> M.block_length sigma.mem obj l
+    | Loc l -> M.block_length sigma obj l
     | Ref x -> noref ~op:"block-length of" x
     | Val(m,x,_) ->
       begin match Ctypes.object_of (vtype m x) with
@@ -795,21 +528,21 @@ struct
   let load sigma obj = function
     | Ref x ->
       begin match V.param x with
-        | ByRef       -> Sigs.Loc(Val(CREF,x,[]))
-        | InContext n -> Sigs.Loc(Val(CTXT n,x,[]))
-        | InArray n   -> Sigs.Loc(Val(CARR n,x,[]))
+        | ByRef       -> Memory.Loc(Val(CREF,x,[]))
+        | InContext n -> Memory.Loc(Val(CTXT n,x,[]))
+        | InArray n   -> Memory.Loc(Val(CARR n,x,[]))
         | NotUsed | ByAddr | ByValue | ByShift -> assert false
       end
     | Val((CREF|CVAL),x,ofs) ->
-      Sigs.Val(access (get_term sigma x) ofs)
+      Memory.Val(access (SVAR.value sigma x) ofs)
     | Loc l ->
       Cvalues.map_value
         (fun l -> Loc l)
-        (M.load sigma.mem obj l)
+        (M.load sigma obj l)
     | Val((CTXT _|CARR _|HEAP) as m,x,ofs) ->
       Cvalues.map_value
         (fun l -> Loc l)
-        (M.load sigma.mem obj (mloc_of_path m x ofs))
+        (M.load sigma obj (mloc_of_path m x ofs))
 
   let load_init sigma obj = function
     | Ref _ ->
@@ -817,23 +550,23 @@ struct
     | Val((CREF|CVAL),x,_) when Cvalues.always_initialized x ->
       Cvalues.initialized_obj obj
     | Val((CREF|CVAL),x,ofs) ->
-      access_init (get_init_term sigma x) ofs
+      access_init (SINIT.value sigma x) ofs
     | Loc l ->
-      M.load_init sigma.mem obj l
+      M.load_init sigma obj l
     | Val((CTXT _|CARR _|HEAP) as m,x,ofs) ->
-      M.load_init sigma.mem obj (mloc_of_path m x ofs)
+      M.load_init sigma obj (mloc_of_path m x ofs)
 
   (* -------------------------------------------------------------------------- *)
   (* ---  Memory Store                                                      --- *)
   (* -------------------------------------------------------------------------- *)
 
   let memvar_stored kind seq x ofs v =
-    let get_term, update = match kind with
-      | KValue -> get_term, update
-      | KInit -> get_init_term, update_init
+    let read, update = match kind with
+      | KValue -> SVAR.value, update
+      | KInit -> SINIT.value, update_init
     in
-    let v1 = get_term seq.pre x in
-    let v2 = get_term seq.post x in
+    let v1 = read seq.pre x in
+    let v2 = read seq.post x in
     Set( v2 , update v1 ofs v )
 
   let gen_stored kind seq obj l v =
@@ -843,17 +576,17 @@ struct
     | Val((CREF|CVAL),x,ofs) ->
       [memvar_stored kind seq x ofs v]
     | Val((CTXT _|CARR _|HEAP) as m,x,ofs) ->
-      mstored (mseq_of_seq seq) obj (mloc_of_path m x ofs) v
+      mstored seq obj (mloc_of_path m x ofs) v
     | Loc l ->
-      mstored (mseq_of_seq seq) obj l v
+      mstored seq obj l v
 
   let stored = gen_stored KValue
   let stored_init = gen_stored KInit
 
   let copied seq obj l1 l2 =
     let v = match load seq.pre obj l2 with
-      | Sigs.Val r -> r
-      | Sigs.Loc l -> pointer_val l
+      | Memory.Val r -> r
+      | Memory.Loc l -> pointer_val l
     in stored seq obj l1 v
 
   let copied_init seq obj l1 l2 =
@@ -1012,7 +745,7 @@ struct
       | RD | OBJ -> p_true
     else
       match mem with
-      | CVAL | HEAP -> p_bool (ALLOC.value sigma.alloc x)
+      | CVAL | HEAP -> p_bool (SALLOC.value sigma x)
       | CREF | CTXT Valid | CARR Valid -> p_true
       | CTXT Nullable | CARR Nullable ->
         p_not @@ M.is_null (mloc_of_path mem x [])
@@ -1035,19 +768,19 @@ struct
     | Rloc(obj,l) ->
       begin match l with
         | Ref _ -> p_true
-        | Loc l -> M.valid sigma.mem acs (Rloc(obj,l))
+        | Loc l -> M.valid sigma acs (Rloc(obj,l))
         | Val(m,x,p) ->
           try valid_offset_path sigma acs m x p
           with ShiftMismatch ->
             if is_heap_allocated m then
-              M.valid sigma.mem acs (Rloc(obj,mloc_of_loc l))
+              M.valid sigma acs (Rloc(obj,mloc_of_loc l))
             else
               shift_mismatch l
       end
     | Rrange(l,elt,a,b) ->
       begin match l with
         | Ref x -> noref ~op:"valid sub-range of" x
-        | Loc l -> M.valid sigma.mem acs (Rrange(l,elt,a,b))
+        | Loc l -> M.valid sigma acs (Rrange(l,elt,a,b))
         | Val(m,x,p) ->
           match a,b with
           | Some ka,Some kb ->
@@ -1058,7 +791,7 @@ struct
               with ShiftMismatch ->
                 if is_heap_allocated m then
                   let l = mloc_of_loc l in
-                  M.valid sigma.mem acs (Rrange(l,elt,a,b))
+                  M.valid sigma acs (Rrange(l,elt,a,b))
                 else shift_mismatch l
             end
           | _ ->
@@ -1085,19 +818,19 @@ struct
     | Rloc(obj,l) ->
       begin match l with
         | Ref _ -> p_false
-        | Loc l -> M.invalid sigma.mem (Rloc(obj,l))
+        | Loc l -> M.invalid sigma (Rloc(obj,l))
         | Val(m,x,p) ->
           try invalid_offset_path sigma m x p
           with ShiftMismatch ->
             if is_heap_allocated m then
-              M.invalid sigma.mem (Rloc(obj,mloc_of_loc l))
+              M.invalid sigma (Rloc(obj,mloc_of_loc l))
             else
               shift_mismatch l
       end
     | Rrange(l,elt,a,b) ->
       begin match l with
         | Ref x -> noref ~op:"invalid sub-range of" x
-        | Loc l -> M.invalid sigma.mem (Rrange(l,elt,a,b))
+        | Loc l -> M.invalid sigma (Rrange(l,elt,a,b))
         | Val(m,x,p) ->
           match a,b with
           | Some ka,Some kb ->
@@ -1108,7 +841,7 @@ struct
               with ShiftMismatch ->
                 if is_heap_allocated m then
                   let l = mloc_of_loc l in
-                  M.invalid sigma.mem (Rrange(l,elt,a,b))
+                  M.invalid sigma (Rrange(l,elt,a,b))
                 else shift_mismatch l
             end
           | _ ->
@@ -1123,7 +856,7 @@ struct
   let rec initialized_loc sigma obj x ofs =
     match obj with
     | C_int _ | C_float _ | C_pointer _ ->
-      p_bool (access_init (get_init_term sigma x) ofs)
+      p_bool (access_init (SINIT.value sigma x) ofs)
     | C_array { arr_flat=flat ; arr_element = te } ->
       let size = match flat with
         | None -> unsized_array ()
@@ -1132,7 +865,7 @@ struct
       initialized_range sigma elt x ofs (e_int 0) (e_int (size-1))
     | C_comp { cfields = None } ->
       Lang.F.p_equal
-        (access_init (get_init_term sigma x) ofs)
+        (access_init (SINIT.value sigma x) ofs)
         (Cvalues.initialized_obj obj)
     | C_comp { cfields = Some fields } ->
       let init_field fd =
@@ -1155,10 +888,10 @@ struct
     | Rloc(obj,l) ->
       begin match l with
         | Ref _ -> p_true
-        | Loc l -> M.initialized sigma.mem (Rloc(obj,l))
+        | Loc l -> M.initialized sigma (Rloc(obj,l))
         | Val(m,x,p) ->
           if is_heap_allocated m then
-            M.initialized sigma.mem (Rloc(obj,mloc_of_loc l))
+            M.initialized sigma (Rloc(obj,mloc_of_loc l))
           else if Cvalues.always_initialized x then
             try valid_offset RD (vobject m x) p
             with ShiftMismatch -> shift_mismatch l
@@ -1168,12 +901,12 @@ struct
     | Rrange(l,elt, Some a, Some b) ->
       begin match l with
         | Ref _ -> p_true
-        | Loc l -> M.initialized sigma.mem (Rrange(l,elt,Some a, Some b))
+        | Loc l -> M.initialized sigma (Rrange(l,elt,Some a, Some b))
         | Val(m,x,p) ->
           try
             if is_heap_allocated m then
               let l = mloc_of_loc l in
-              M.initialized sigma.mem (Rrange(l,elt,Some a, Some b))
+              M.initialized sigma (Rrange(l,elt,Some a, Some b))
             else
               let rec normalize obj = function
                 | [] -> [], a, b
@@ -1220,18 +953,21 @@ struct
 
   let frame sigma =
     let hs = ref [] in
-    SIGMA.iter
-      begin fun x chunk ->
-        (if (x.vglob || x.vformal) then
-           let t = VAR.typ_of_chunk x in
-           let v = e_var chunk in
-           let h = forall_pointers (M.global sigma.mem) v t in
-           if not (F.eqp h F.p_true) then hs := h :: !hs ) ;
-        (if x.vglob then
-           let v = e_var chunk in
-           hs := Cvalues.has_ctype x.vtype v :: !hs ) ;
-      end sigma.vars ;
-    !hs @ M.frame sigma.mem
+    Sigma.iter
+      begin fun chunk value ->
+        match Sigma.ckind chunk with
+        | SVAR.Mu x when not @@ is_ref x ->
+          (if (x.vglob || x.vformal) then
+             let t = VAR.typ_of_chunk x in
+             let v = e_var value in
+             let h = forall_pointers (M.global sigma) v t in
+             if not (F.eqp h F.p_true) then hs := h :: !hs ) ;
+          (if x.vglob then
+             let v = e_var value in
+             hs := Cvalues.has_ctype x.vtype v :: !hs ) ;
+        | _ -> ()
+      end sigma ;
+    !hs @ M.frame sigma
 
   (* -------------------------------------------------------------------------- *)
   (* ---  Scope                                                             --- *)
@@ -1248,25 +984,24 @@ struct
 
   let alloc sigma xs =
     let xm = List.filter is_mem xs in
-    let mem = M.alloc sigma.mem xm in
+    let sigma = M.alloc sigma xm in
     let xv = List.filter is_mvar_alloc xs in
-    let domain = TALLOC.Set.of_list xv in
-    let alloc = ALLOC.havoc sigma.alloc domain in
-    { sigma with alloc ; mem }
+    let domain = Sigma.Domain.of_list @@ List.map SALLOC.chunk xv in
+    Sigma.havoc sigma domain
 
   let scope_vars seq scope xs =
     let xs = List.filter is_mvar_alloc xs in
     if xs = [] then []
     else
-      let t_in = seq.pre.alloc in
-      let t_out = seq.post.alloc in
+      let t_in = seq.pre in
+      let t_out = seq.post in
       let v_in  = match scope with Enter -> e_false | Leave -> e_true in
       let v_out = match scope with Enter -> e_true | Leave -> e_false in
       List.map
         (fun x ->
            F.p_and
-             (F.p_equal (ALLOC.value t_in x) v_in)
-             (F.p_equal (ALLOC.value t_out x) v_out)
+             (F.p_equal (SALLOC.value t_in x) v_in)
+             (F.p_equal (SALLOC.value t_out x) v_out)
         ) xs
 
   let scope_init seq scope xs =
@@ -1278,7 +1013,7 @@ struct
         then None
         else
           let i = Cvalues.uninitialized_obj (Ctypes.object_of v.vtype) in
-          Some (Lang.F.p_equal (access_init (get_init_term seq.post v) []) i)
+          Some (Lang.F.p_equal (access_init (SINIT.value seq.post v) []) i)
       in
       List.filter_map init_status xs
 
@@ -1296,14 +1031,13 @@ struct
 
   let scope seq scope xs =
     let xm = List.filter is_mem xs in
-    let smem = { pre = seq.pre.mem ; post = seq.post.mem } in
-    let hmem = M.scope smem scope xm in
+    let hmem = M.scope seq scope xm in
     let hvars = scope_vars seq scope xs in
     let hinit = scope_init seq scope xs in
     let hcontext = List.filter_map nullable_scope xs in
     hvars @ hinit @ hmem @ hcontext
 
-  let global sigma p = M.global sigma.mem p
+  let global = M.global
 
   (* -------------------------------------------------------------------------- *)
   (* ---  Havoc along a ranged-path                                        --- *)
@@ -1358,9 +1092,9 @@ struct
           assigned_path kind (eqk :: hs) xs ys ae be ofs
 
   let assigned_genset s xs mem x ofs p =
-    let valid = valid_offset_path s.post Sigs.RW mem x ofs in
-    let a = get_term s.pre x in
-    let b = get_term s.post x in
+    let valid = valid_offset_path s.post Memory.RW mem x ofs in
+    let a = SVAR.value s.pre x in
+    let b = SVAR.value s.post x in
     let a_ofs = access a ofs in
     let b_ofs = access b ofs in
     let p_sloc = p_forall xs (p_hyps [valid;p_not p] (p_equal a_ofs b_ofs)) in
@@ -1371,9 +1105,9 @@ struct
   let monotonic_initialized_genset s xs mem x ofs p =
     if always_init x then [Assert p_true]
     else
-      let valid = valid_offset_path s.post Sigs.RW mem x ofs in
-      let a = get_init_term s.pre x in
-      let b = get_init_term s.post x in
+      let valid = valid_offset_path s.post Memory.RW mem x ofs in
+      let a = SINIT.value s.pre x in
+      let b = SINIT.value s.post x in
       let a_ofs = access_init a ofs in
       let b_ofs = access_init b ofs in
       let not_p = p_forall xs (p_hyps [valid;p_not p] (p_equal a_ofs b_ofs)) in
@@ -1399,8 +1133,8 @@ struct
 
       | C_int _ | C_float _ | C_pointer _ ->
         p_imply
-          (p_bool (access_init (get_init_term seq.pre x) ofs))
-          (p_bool (access_init (get_init_term seq.post x) ofs))
+          (p_bool (access_init (SINIT.value seq.pre x) ofs))
+          (p_bool (access_init (SINIT.value seq.post x) ofs))
 
       | C_array { arr_flat=flat ; arr_element=t } ->
         let size = match flat with
@@ -1430,9 +1164,9 @@ struct
         [ memvar_stored KInit seq x ofs (e_var init) ]
       end
     | Val((HEAP|CTXT _|CARR _) as m,x,ofs) ->
-      M.assigned (mseq_of_seq seq) obj (Sloc (mloc_of_path m x ofs))
+      M.assigned seq obj (Sloc (mloc_of_path m x ofs))
     | Loc l ->
-      M.assigned (mseq_of_seq seq) obj (Sloc l)
+      M.assigned seq obj (Sloc l)
 
   let assigned_array seq obj l elt n =
     match l with
@@ -1452,17 +1186,17 @@ struct
 
     | Val((HEAP|CTXT _|CARR _) as m,x,ofs) ->
       let l = mloc_of_path m x ofs in
-      M.assigned (mseq_of_seq seq) obj (Sarray(l,elt,n))
+      M.assigned seq obj (Sarray(l,elt,n))
     | Loc l ->
-      M.assigned (mseq_of_seq seq) obj (Sarray(l,elt,n))
+      M.assigned seq obj (Sarray(l,elt,n))
 
   let assigned_range seq obj l elt a b =
     match l with
     | Ref x -> noref ~op:"assigns to" x
     | Loc l ->
-      M.assigned (mseq_of_seq seq) obj (Srange(l,elt,a,b))
+      M.assigned seq obj (Srange(l,elt,a,b))
     | Val((HEAP|CTXT _|CARR _) as m,x,ofs) ->
-      M.assigned (mseq_of_seq seq) obj (Srange(mloc_of_path m x ofs,elt,a,b))
+      M.assigned seq obj (Srange(mloc_of_path m x ofs,elt,a,b))
     | Val((CVAL|CREF) as m,x,ofs) ->
       let k = Lang.freshvar ~basename:"k" Qed.Logic.Int in
       let p = Vset.in_range (e_var k) a b in
@@ -1474,9 +1208,9 @@ struct
     match l with
     | Ref x -> noref ~op:"assigns to" x
     | Loc l ->
-      M.assigned (mseq_of_seq seq) obj (Sdescr(xs,l,p))
+      M.assigned seq obj (Sdescr(xs,l,p))
     | Val((HEAP|CTXT _|CARR _) as m,x,ofs) ->
-      M.assigned (mseq_of_seq seq) obj (Sdescr(xs,mloc_of_path m x ofs,p))
+      M.assigned seq obj (Sdescr(xs,mloc_of_path m x ofs,p))
     | Val((CVAL|CREF) as m,x,ofs) ->
       (* (monotonic_initialized_genset seq xs m x ofs p) @ *)
       (assigned_genset seq xs m x ofs p)
@@ -1624,20 +1358,22 @@ struct
     match l with
     | Ref x | Val((CVAL|CREF),x,_) ->
       let init =
-        if not @@ Cvalues.always_initialized x then [ Init x ] else []
+        if not @@ Cvalues.always_initialized x then [ SINIT.chunk x ] else []
       in
-      Heap.Set.of_list ((Var x) :: init)
+      Sigma.Domain.of_list (SVAR.chunk x :: init)
     | Loc _ | Val((CTXT _|CARR _|HEAP),_,_) ->
-      M.Heap.Set.fold
-        (fun m s -> Heap.Set.add (Mem m) s)
-        (M.domain obj (mloc_of_loc l)) Heap.Set.empty
+      M.domain obj (mloc_of_loc l)
 
   let is_well_formed sigma =
     let cstrs = ref [] in
-    SIGMA.iter
-      (fun v c -> cstrs := Cvalues.has_ctype v.vtype (e_var c) :: !cstrs)
-      sigma.vars ;
-    p_conj ((M.is_well_formed sigma.mem) :: !cstrs)
+    Sigma.iter
+      (fun chunk value ->
+         match Sigma.ckind chunk with
+         | SVAR.Mu x when not @@ is_ref x ->
+           cstrs := Cvalues.has_ctype x.vtype (e_var value) :: !cstrs
+         | _ -> ())
+      sigma ;
+    p_conj (M.is_well_formed sigma :: !cstrs)
 
   (* -------------------------------------------------------------------------- *)
 
diff --git a/src/plugins/wp/MemVar.mli b/src/plugins/wp/MemVar.mli
index 468ebe8a2470ac3368907ee14f276c06e3d6ebe1..3c88cb04ddf1186a47b626ab96e0635046b68f8d 100644
--- a/src/plugins/wp/MemVar.mli
+++ b/src/plugins/wp/MemVar.mli
@@ -42,4 +42,4 @@ module Raw : VarUsage
 module Static : VarUsage
 
 (** Create a mixed Hoare Memory Model from VarUsage instance. *)
-module Make(_ : VarUsage)(_ : Sigs.Model) : Sigs.Model
+module Make(_ : VarUsage)(_ : Memory.Model) : Memory.Model
diff --git a/src/plugins/wp/MemZeroAlias.ml b/src/plugins/wp/MemZeroAlias.ml
index efc890722c7a22e6b2e395e01f6eb9458aacddf0..735fa9ecaea33973494eb51b238b0bbb499a8159 100644
--- a/src/plugins/wp/MemZeroAlias.ml
+++ b/src/plugins/wp/MemZeroAlias.ml
@@ -28,7 +28,7 @@ open Cil_types
 open Cil_datatype
 open Lang
 open Lang.F
-open Sigs
+open Memory
 
 module Logic = Qed.Logic
 
@@ -115,11 +115,12 @@ struct
     let te = Lang.tau_of_object (object_of_rpath x (List.rev p)) in
     dim_of_path te p
   let basename_of_chunk (x,_) = LogicUsage.basename x
+  let is_init _ = false
+  let is_primary _ = true
   let is_framed (x,p) = not x.vglob && p = []
 end
 
-module Heap = Qed.Collection.Make(Chunk)
-module Sigma = Sigma.Make(Chunk)(Heap)
+module State = Sigma.Make(Chunk)
 
 type loc =
   | Null
@@ -128,8 +129,6 @@ type loc =
   | Array of loc * F.term
   | Field of loc * fieldinfo
 
-type sigma = Sigma.t
-type domain = Sigma.domain
 type segment = loc rloc
 
 let rec pretty fmt = function
@@ -180,14 +179,14 @@ let rec walk ps ks = function
 let access l = walk [] [] l
 
 let domain _obj l =
-  try Heap.Set.singleton (fst (access l))
-  with _ -> Heap.Set.empty
+  try State.singleton (fst (access l))
+  with _ -> Sigma.empty
 
 let is_well_formed _s = p_true
 
 let value sigma l =
   let m,ks = access l in
-  let x = Sigma.get sigma m in
+  let x = State.get sigma m in
   List.fold_left F.e_get (e_var x) ks
 
 let rec update a ks v =
@@ -195,7 +194,7 @@ let rec update a ks v =
   | [] -> v
   | k::ks -> F.e_set a k (update (F.e_get a k) ks v)
 
-let set s m ks v = if ks = [] then v else update (e_var (Sigma.get s m)) ks v
+let set s m ks v = if ks = [] then v else update (State.value s m) ks v
 
 let load sigma obj l =
   if Ctypes.is_pointer obj then Loc (Star l) else Val(value sigma l)
@@ -204,7 +203,7 @@ let load_init _sigma _obj _l = Warning.error ~source "Mem0Alias: No initialized"
 
 let stored seq _obj l e =
   let m,ks = access l in
-  let x = F.e_var (Sigma.get seq.post m) in
+  let x = State.value seq.post m in
   [Set( x , set seq.pre m ks e )]
 
 let stored_init _seq _obj _l _e = Warning.error ~source "Mem0Alias: No initialized"
@@ -216,18 +215,6 @@ let copied_init _seq _obj _a _b = Warning.error ~source "Mem0Alias: No initializ
 
 let assigned _s _obj _sloc = []
 
-type state = Chunk.t Tmap.t
-let state (s:sigma) =
-  let m = ref Tmap.empty in
-  Sigma.iter (fun c x -> m := Tmap.add (F.e_var x) c !m) s ; !m
-
-let imval c = Sigs.Mchunk (Pretty_utils.to_string Chunk.pretty c, KValue)
-let iter f s = Tmap.iter (fun v c -> f (imval c) v) s
-let lookup (s : state) (e : Lang.F.term) = imval (F.Tmap.find e s)
-let apply f s =
-  let m = ref Tmap.empty in
-  Tmap.iter (fun e c -> m := Tmap.add (f e) c !m) s ; !m
-
 let rec ipath lv = function
   | [] -> lv
   | S::w -> ipath (Mval lv,[]) w
@@ -237,28 +224,40 @@ let rec ipath lv = function
     ipath (host, path @ [Mfield f]) w
 let ilval (x,p) = ipath (Mvar x,[]) p
 
+let lookup (s : Sigma.state) (e : F.term) =
+  match Sigma.ckind @@ Lang.F.Tmap.find e s with
+  | State.Mu lv -> Mlval (ilval lv)
+  | _ -> raise Not_found
+
+module Hmap = Sigma.Heap.Map
+
 let heap domain state =
   Tmap.fold
     (fun m c w ->
        if Vars.intersect (F.vars m) domain
-       then Heap.Map.add c m w else w
-    ) state Heap.Map.empty
+       then Hmap.add c m w else w
+    ) state Hmap.empty
 
 let updates seq domain =
   let pre = heap domain seq.pre in
   let post = heap domain seq.post in
   let pool = ref Bag.empty in
-  Heap.Map.iter2
+  Hmap.iter2
     (fun c v1 v2 ->
-       try
-         match v1,v2 with
-         | _,None -> ()
-         | None,Some v ->
-           pool := Bag.add (Mstore(ilval c,v)) !pool
-         | Some v1,Some v2 ->
-           if v2 != v1 then
-             pool := Bag.add (Mstore (ilval c,v2)) !pool
-       with Not_found -> ()
+       match Sigma.ckind c with
+       | State.Mu m ->
+         begin
+           try
+             match v1,v2 with
+             | _,None -> ()
+             | None,Some v ->
+               pool := Bag.add (Mstore(ilval m,v)) !pool
+             | Some v1,Some v2 ->
+               if v2 != v1 then
+                 pool := Bag.add (Mstore (ilval m,v2)) !pool
+           with Not_found -> ()
+         end
+       | _ -> ()
     ) pre post ;
   !pool
 
diff --git a/src/plugins/wp/MemZeroAlias.mli b/src/plugins/wp/MemZeroAlias.mli
index 8523b1df8cafd65934b46f4554c4c8e850896068..4fc1b752cb1eee50883707fc41a1604fa6e3ea75 100644
--- a/src/plugins/wp/MemZeroAlias.mli
+++ b/src/plugins/wp/MemZeroAlias.mli
@@ -24,4 +24,4 @@
 (* --- Empty Memory Model                                                 --- *)
 (* -------------------------------------------------------------------------- *)
 
-include Sigs.Model
+include Memory.Model
diff --git a/src/plugins/wp/Sigs.ml b/src/plugins/wp/Memory.ml
similarity index 76%
rename from src/plugins/wp/Sigs.ml
rename to src/plugins/wp/Memory.ml
index da6c1eb517d8301ac608d05618628a2f5121e394..3398d7f045d59a690da60f5eaf8a95a3b173980a 100644
--- a/src/plugins/wp/Sigs.ml
+++ b/src/plugins/wp/Memory.ml
@@ -26,9 +26,9 @@
 
 open Cil_types
 open Ctypes
-open Lang
 open Lang.F
 open Interpreted_automata
+open Sigma
 
 (* -------------------------------------------------------------------------- *)
 (** {1 General Definitions} *)
@@ -116,10 +116,10 @@ and s_offset = Mfield of fieldinfo | Mindex of term
 
 (** Reversed abstract value *)
 type mval =
-  | Mterm (** Not a state-related value *)
   | Maddr of s_lval (** The value is the address of an l-value in current memory *)
-  | Mlval of s_lval * datakind (** The value is the value of an l-value in current memory *)
-  | Mchunk of string * datakind (** The value is an abstract memory chunk (description) *)
+  | Mlval of s_lval (** The value is the value of an l-value in current memory *)
+  | Minit of s_lval (** The value is the init state of an l-value in current memory *)
+  | Mchunk of Chunk.t (** The value is an abstract memory chunk (description) *)
 
 (** Reversed update *)
 type update = Mstore of s_lval * term
@@ -129,169 +129,6 @@ type update = Mstore of s_lval * term
 (** {1 Memory Models} *)
 (* -------------------------------------------------------------------------- *)
 
-module type Type =
-sig
-  type t
-  val hash : t -> int
-  val equal : t -> t -> bool
-  val compare : t -> t -> int
-  val pretty : Format.formatter -> t -> unit
-end
-
-(** Memory Chunks.
-
-    The concrete memory is partionned into a vector of abstract data.
-    Each component of the partition is called a {i memory chunk} and
-    holds an abstract representation of some part of the memory.
-
-    Remark: memory chunks are not required to be independant from each other,
-    provided the memory model implementation is consistent with the chosen
-    representation. Conversely, a given object might be represented by
-    several memory chunks. See {!Model.domain}.
-
-*)
-module type Chunk =
-sig
-
-  type t
-  val self : string
-  (** Chunk names, for pretty-printing. *)
-
-  include Type with type t := t
-
-  val tau_of_chunk : t -> tau
-  (** The type of data hold in a chunk. *)
-
-  val basename_of_chunk : t -> string
-  (** Used when generating fresh variables for a chunk. *)
-
-  val is_framed : t -> bool
-  (** Whether the chunk is local to a function call.
-
-      Means the chunk is separated from anyother call side-effects.
-      If [true], entails that a function assigning everything can not modify
-      the chunk. Only used for optimisation, it would be safe to always
-      return [false]. *)
-
-end
-
-(** Memory Environments.
-
-    Represents the content of the memory, {i via} a vector of logic
-    variables for each memory chunk.
-*)
-module type Sigma =
-sig
-
-  type chunk
-  (** The type of memory chunks. *)
-
-  module Chunk : Qed.Collection.S with type t = chunk
-
-  (** Memory footprint. *)
-  type domain = Chunk.Set.t
-
-  (** Environment assigning logic variables to chunk.
-
-      Memory chunk variables are assigned lazily. Hence, the vector is
-      empty unless a chunk is accessed. Pay attention to this
-      when you merge or havoc chunks (using memcpy).
-
-      New chunks are generated from the context pool of {!Lang.freshvar}.
-  *)
-  type t
-
-  val pretty : Format.formatter -> t -> unit
-  (** For debugging purpose *)
-
-  val create : unit -> t
-  (** Initially empty environment. *)
-
-  val mem : t -> chunk -> bool
-  (** Whether a chunk has been assigned. *)
-
-  val get : t -> chunk -> var
-  (** Lazily get the variable for a chunk. *)
-
-  val value : t -> chunk -> term
-  (** Same as [Lang.F.e_var] of [get]. *)
-
-  val copy : t -> t (** Duplicate the environment. Fresh chunks in the copy
-                        are {i not} duplicated into the source environment. *)
-
-  val join : t -> t -> Passive.t
-  (** Make two environment pairwise equal {i via} the passive form.
-
-      Missing chunks in one environment are added with the corresponding
-      variable of the other environment. When both environments don't agree
-      on a chunk, their variables are added to the passive form. *)
-
-  val assigned : pre:t -> post:t -> domain -> pred Bag.t
-  (** Make chunks equal outside of some domain.
-
-      This is similar to [join], but outside the given footprint of an
-      assigns clause. Although, the function returns the equality
-      predicates instead of a passive form.
-
-      Like in [join], missing chunks are reported from one side to the
-      other one, and common chunks are added to the equality bag. *)
-
-  val choose : t -> t -> t
-  (** Make the union of each sigma, choosing the minimal variable
-      in case of conflict.
-      Both initial environments are kept unchanged. *)
-
-  val merge : t -> t -> t * Passive.t * Passive.t
-  (** Make the union of each sigma, choosing a {i new} variable for
-      each conflict, and returns the corresponding joins.
-      Both initial environments are kept unchanged. *)
-
-  val merge_list : t list -> t * Passive.t list
-  (** Same than {!merge} but for a list of sigmas. Much more efficient
-      than folding merge step by step. *)
-
-  val iter : (chunk -> var -> unit) -> t -> unit
-  (** Iterates over the chunks and associated variables already
-      accessed so far in the environment. *)
-
-  val iter2 : (chunk -> var option -> var option -> unit) -> t -> t -> unit
-  (** Same as [iter] for both environments. *)
-
-  val havoc_chunk : t -> chunk -> t
-  (** Generate a new fresh variable for the given chunk. *)
-
-  val havoc : t -> domain -> t
-  (** All the chunks in the provided footprint are generated and made fresh.
-
-      Existing chunk variables {i outside} the footprint are copied into the new
-      environment. The original environement itself is kept unchanged. More
-      efficient than iterating [havoc_chunk] over the footprint.
-  *)
-
-  val havoc_any : call:bool -> t -> t
-  (** All the chunks are made fresh. As an optimisation,
-      when [~call:true] is set, only non-local chunks are made fresh.
-      Local chunks are those for which [Chunk.is_frame] returns [true]. *)
-
-  val remove_chunks : t -> domain -> t
-  (** Return a copy of the environment where chunks in the footprint
-      have been removed. Keep the original environment unchanged. *)
-
-  val domain : t -> domain
-  (** Footprint of a memory environment.
-      That is, the set of accessed chunks so far in the environment. *)
-
-  val union : domain -> domain -> domain
-  (** Same as [Chunk.Set.union] *)
-
-  val empty : domain
-  (** Same as [Chunk.Set.empty] *)
-
-  val writes : t sequence -> domain
-  (** [writes s] indicates which chunks are new in [s.post] compared
-      to [s.pre]. *)
-end
-
 (** Memory Models. *)
 module type Model =
 sig
@@ -315,34 +152,13 @@ sig
       This function typically adds new elements to the partition received
       in input (that can be empty). *)
 
-  module Chunk : Chunk
-  (** Memory model chunks. *)
-
-  module Heap : Qed.Collection.S
-    with type t = Chunk.t
-  (** Chunks Sets and Maps. *)
-
-  module Sigma : Sigma
-    with type chunk = Chunk.t
-     and module Chunk = Heap
-  (** Model Environments. *)
-
   type loc
   (** Representation of the memory location in the model. *)
 
-  type chunk = Chunk.t
-  type sigma = Sigma.t
-  type domain = Sigma.domain
   type segment = loc rloc
 
   (** {2 Reversing the Model} *)
 
-  type state
-  (** Internal (private) memory state description for later reversing the model. *)
-
-  (** Returns a memory state description from a memory environement. *)
-  val state : sigma -> state
-
   (** Try to interpret a term as an in-memory operation
       located at this program point. Only best-effort
       shall be performed, otherwise return [Mvalue].
@@ -355,17 +171,11 @@ sig
 
   (** Try to interpret a sequence of states into updates.
 
-      The result shall be exhaustive with respect to values that are printed as [Sigs.mval]
+      The result shall be exhaustive with respect to values that are printed as [Memory.mval]
       values at [post] label {i via} the [lookup] function.
       Otherwise, those values would not be pretty-printed to the user. *)
   val updates : state sequence -> Vars.t -> update Bag.t
 
-  (** Propagate a sequent substitution inside the memory state. *)
-  val apply : (term -> term) -> state -> state
-
-  (** Debug *)
-  val iter : (mval -> term -> unit) -> state -> unit
-
   val pretty : Format.formatter -> loc -> unit
   (** pretty printing of memory location *)
 
@@ -558,7 +368,6 @@ sig
   type loc = M.loc
   type nonrec value = loc value
   type nonrec result = loc result
-  type sigma = M.Sigma.t
 
   val pp_value : Format.formatter -> value -> unit
 
@@ -638,28 +447,27 @@ sig
     value option ->
     pred
 
-  val unchanged : M.sigma -> M.sigma -> varinfo -> pred
+  val unchanged : sigma -> sigma -> varinfo -> pred
   (** Express that a given variable has the same value in two memory states. *)
 
   type warned_hyp = Warning.Set.t * (pred * pred)
 
-  val init :
-    sigma:M.sigma -> varinfo -> init option -> warned_hyp list
-    (** Express that some variable has some initial value at the
-        given memory state. The first predicate states the value,
-        the second, the initialization status.
-
-        Note: we DO NOT merge values and initialization status
-        hypotheses as the factorization performed by Qed can make
-        predicates too hard to simplify later.
-
-        Remark: [None] initializer are interpreted as zeroes. This is consistent
-        with the [init option] associated with global variables in CIL,
-        for which the default initializer are zeroes. This function is called
-        for global initializers and local initializers ([Cil.Local_init]).
-        It is not called for local variables without initializers as they do not
-        have a [Cil.init option].
-    *)
+  val init : sigma:sigma -> varinfo -> init option -> warned_hyp list
+  (** Express that some variable has some initial value at the
+      given memory state. The first predicate states the value,
+      the second, the initialization status.
+
+      Note: we DO NOT merge values and initialization status
+      hypotheses as the factorization performed by Qed can make
+      predicates too hard to simplify later.
+
+      Remark: [None] initializer are interpreted as zeroes. This is consistent
+      with the [init option] associated with global variables in CIL,
+      for which the default initializer are zeroes. This function is called
+      for global initializers and local initializers ([Cil.Local_init]).
+      It is not called for local variables without initializers as they do not
+      have a [Cil.init option].
+  *)
 
 end
 
@@ -674,7 +482,6 @@ sig
   type nonrec logic  = M.loc logic
   type nonrec region = M.loc region
   type nonrec result = M.loc result
-  type sigma = M.Sigma.t
 
   (** {2 Frames}
 
diff --git a/src/plugins/wp/Mstate.ml b/src/plugins/wp/Mstate.ml
index 125f6139e5ec4a9de88bd0abf5f12886f02d0330..61b71271b88f382739f4de4d8b61bdf2c86a47d8 100644
--- a/src/plugins/wp/Mstate.ml
+++ b/src/plugins/wp/Mstate.ml
@@ -20,25 +20,7 @@
 (*                                                                        *)
 (**************************************************************************)
 
-open Lang.F
-open Sigs
-
-type (_,_) eq = Equal : 'a -> ('a,'a) eq | NotEqual : ('a,'b) eq
-
-module Ident :
-sig
-  type 'a t
-  val create : 'a -> 'a t
-  val get : 'a t -> 'a
-  val eq : 'a t -> 'b t -> ('a,'b) eq
-end =
-struct
-  let k = ref 0
-  type 'a t = int * 'a
-  let get = snd
-  let create s = incr k ; (!k,s)
-  let eq (k,s) (k',_) = if k = k' then Obj.magic (Equal s) else NotEqual
-end
+open Memory
 
 (* -------------------------------------------------------------------------- *)
 (* --- L-Val Utility                                                      --- *)
@@ -68,35 +50,41 @@ let equal a b = a == b || (host_eq (fst a) (fst b) && offset_eq (snd a) (snd b))
 (* --- Memory State Pretty Printing Information                           --- *)
 (* -------------------------------------------------------------------------- *)
 
-type 'a operations = {
-  apply : (term -> term) -> 'a -> 'a ;
-  lookup : 'a -> term -> mval ;
-  updates : 'a sequence -> Vars.t -> update Bag.t ;
-  iter : (mval -> term -> unit) -> 'a -> unit ;
-}
-
-type 'a model = MODEL : 'a operations Ident.t * ('b -> 'a) -> 'b model
+type state = STATE of {
+    model : (module Memory.Model) ;
+    state : Sigma.state ;
+  }
 
-let create (type s) (module M : Sigs.Model with type Sigma.t = s) =
-  let op = {
-    apply = M.apply ;
-    lookup = M.lookup ;
-    updates = M.updates ;
-    iter = M.iter ;
-  } in
-  MODEL( Ident.create op , M.state )
+let create (module M: Memory.Model) sigma =
+  STATE { model = (module M) ; state = Sigma.state sigma }
 
-type state = STATE : 'a operations Ident.t * 'a -> state
+let apply f = function STATE { model ; state } ->
+  let module M = (val model) in
+  let state = Sigma.apply f state in
+  STATE { model ; state }
 
-let state model sigma = match model with MODEL(op,state) -> STATE(op,state sigma)
+let lookup s e =
+  match s with STATE { model ; state } ->
+    let module M = (val model) in
+    try M.lookup state e
+    with Not_found ->
+      Mchunk (Lang.F.Tmap.find e state)
 
-let iter f = function STATE(m,s) -> (Ident.get m).iter f s
-let apply f = function STATE(m,s) -> STATE(m,(Ident.get m).apply f s)
-let lookup s e = match s with STATE(m,s) -> (Ident.get m).lookup s e
 let updates seq vars =
-  match seq.pre , seq.post with STATE(p,u) , STATE(q,v) ->
-  match Ident.eq p q with
-  | Equal s -> s.updates { pre = u ; post = v } vars
-  | NotEqual -> Bag.empty (* assert false *)
+  match seq.pre, seq.post with
+    STATE { model ; state = pre },
+    STATE { state = post } ->
+    let module M = (val model) in
+    M.updates { pre ; post } vars
+
+let iter f s =
+  match s with STATE { model ; state } ->
+    let module M = (val model) in
+    Lang.F.Tmap.iter
+      (fun value chunk ->
+         match M.lookup state value with
+         | exception Not_found -> f (Mchunk chunk) value
+         | mval -> f mval value
+      ) state
 
 (* -------------------------------------------------------------------------- *)
diff --git a/src/plugins/wp/Mstate.mli b/src/plugins/wp/Mstate.mli
index 6931277b596bd0a88e6e298459368013dd92dc20..64e257078d01c288cfaa7c1866174e562fa32902 100644
--- a/src/plugins/wp/Mstate.mli
+++ b/src/plugins/wp/Mstate.mli
@@ -21,7 +21,8 @@
 (**************************************************************************)
 
 open Lang.F
-open Sigs
+open Memory
+open Sigma
 
 (* -------------------------------------------------------------------------- *)
 (* --- L-Val Utility                                                      --- *)
@@ -35,11 +36,9 @@ val equal : s_lval -> s_lval -> bool
 (* --- Memory State Pretty Printing Information                           --- *)
 (* -------------------------------------------------------------------------- *)
 
-type 'a model
 type state
 
-val create : (module Model with type Sigma.t = 'a) -> 'a model
-val state : 'a model -> 'a -> state
+val create : (module Memory.Model) -> sigma -> state
 
 val lookup : state -> term -> mval
 val apply : (term -> term) -> state -> state
diff --git a/src/plugins/wp/Pcfg.ml b/src/plugins/wp/Pcfg.ml
index 5e6f1884cdb21aea32d6a974e2bfa30f46767eae..99321f369f5822378960b823d8f8d2a460c62082 100644
--- a/src/plugins/wp/Pcfg.ml
+++ b/src/plugins/wp/Pcfg.ml
@@ -23,7 +23,7 @@
 open Cil_types
 open Lang
 open Lang.F
-open Sigs
+open Memory
 
 (* -------------------------------------------------------------------------- *)
 (* --- State Registry                                                     --- *)
@@ -43,7 +43,7 @@ type value =
   | Addr of s_lval
   | Lval of s_lval * label
   | Init of s_lval * label
-  | Chunk of string * label
+  | Chunk of Sigma.chunk * label
 
 module Imap = Datatype.Int.Map
 
@@ -91,12 +91,12 @@ let rec find env e =
 and lookup env e = function
   | [] -> Term
   | lbl :: others ->
-    try match Mstate.lookup lbl.state e with
-      | Sigs.Mterm -> raise Not_found
-      | Sigs.Maddr lv -> Addr lv
-      | Sigs.Mlval (lv, KValue) -> Lval(lv,flag lbl)
-      | Sigs.Mlval (lv, KInit) -> Init(lv,flag lbl)
-      | Sigs.Mchunk (m, _) -> Chunk(m,flag lbl)
+    try
+      match Mstate.lookup lbl.state e with
+      | Memory.Maddr lv -> Addr lv
+      | Memory.Mlval lv -> Lval(lv,flag lbl)
+      | Memory.Minit lv -> Init(lv,flag lbl)
+      | Memory.Mchunk m -> Chunk(m,flag lbl)
     with Not_found -> lookup env e others
 
 let is_ref x k = (k == F.e_zero) && Cil.isPointerType x.vtype
@@ -108,7 +108,7 @@ let is_atomic = function
 let iter f lbl = Mstate.iter f lbl.state
 
 let is_copy env lbl = function
-  | Sigs.Mstore( lv , value ) ->
+  | Memory.Mstore( lv , value ) ->
     begin
       match find env value with
       | Lval(lv0,lbl0) -> lbl0 == lbl && Mstate.equal lv lv0
@@ -218,9 +218,9 @@ class virtual engine =
     method pp_offset fmt fs = List.iter (self#pp_ofs fmt) fs
 
     method pp_host fmt = function
-      | Sigs.Mvar x -> Format.pp_print_string fmt x.vname
-      | Sigs.Mmem p -> self#pp_atom fmt p
-      | Sigs.Mval lv -> self#pp_lval fmt lv
+      | Memory.Mvar x -> Format.pp_print_string fmt x.vname
+      | Memory.Mmem p -> self#pp_atom fmt p
+      | Memory.Mval lv -> self#pp_lval fmt lv
 
     method pp_lval fmt = function
       | Mvar x , [] ->
@@ -251,11 +251,11 @@ class virtual engine =
     method pp_label fmt lbl =
       Format.pp_print_string fmt lbl.name
 
-    method pp_chunk fmt m = Format.fprintf fmt "µ:%s" m
+    method pp_chunk fmt mu = Format.fprintf fmt "µ:%a" Sigma.Chunk.pretty mu
 
   end
 
-open Sigs
+open Memory
 
 let rec lv_iter f (h,ofs) = host_iter f h ; List.iter (ofs_iter f) ofs
 and host_iter f = function Mvar _ -> () | Mmem e -> f e | Mval lv -> lv_iter f lv
diff --git a/src/plugins/wp/Pcfg.mli b/src/plugins/wp/Pcfg.mli
index 5fe8d803eccda6c192f2a9f4862a84cbdbf124b9..2c5a65c217428b1d9a619558eae720ad462f8d0d 100644
--- a/src/plugins/wp/Pcfg.mli
+++ b/src/plugins/wp/Pcfg.mli
@@ -20,7 +20,7 @@
 (*                                                                        *)
 (**************************************************************************)
 
-open Sigs
+open Memory
 open Lang
 open Lang.F
 
@@ -32,19 +32,19 @@ type value =
   | Addr of s_lval
   | Lval of s_lval * label
   | Init of s_lval * label
-  | Chunk of string * label
+  | Chunk of Sigma.chunk * label
 
 val create : unit -> env
 val register : Conditions.sequence -> env
 
 val at : env -> id:int -> label
 val find : env -> F.term -> value
-val updates : env -> label Sigs.sequence -> Vars.t -> Sigs.update Bag.t
+val updates : env -> label Memory.sequence -> Vars.t -> Memory.update Bag.t
 val visible : label -> bool
 val subterms : env -> (F.term -> unit) -> F.term -> bool
 val prev : label -> label list
 val next : label -> label list
-val iter : (Sigs.mval -> term -> unit) -> label -> unit
+val iter : (Memory.mval -> term -> unit) -> label -> unit
 val branching : label -> bool
 
 class virtual engine :
@@ -67,5 +67,5 @@ class virtual engine :
 
     method pp_label : Format.formatter -> label -> unit (** label name *)
 
-    method pp_chunk : Format.formatter -> string -> unit (** chunk name *)
+    method pp_chunk : Format.formatter -> Sigma.chunk -> unit (** chunk name *)
   end
diff --git a/src/plugins/wp/Pcond.ml b/src/plugins/wp/Pcond.ml
index 12608d0e3cc7fd7d9ace31c64d920a08ddb35692..9fabe080cd6e293ba83abeff1bffb777426df2be 100644
--- a/src/plugins/wp/Pcond.ml
+++ b/src/plugins/wp/Pcond.ml
@@ -162,7 +162,7 @@ class state =
 
     method updates seq = Pcfg.updates env seq domain
 
-    method pp_update lbl fmt = function Sigs.Mstore(lv,v) ->
+    method pp_update lbl fmt = function Memory.Mstore(lv,v) ->
       let stack = context in
       context <- AtLabel lbl ;
       Format.fprintf fmt "@[<hov 2>%a =@ %a;@]"
@@ -391,13 +391,13 @@ class seqengine (lang : #state) =
       | Some lbl ->
         let before = match Pcfg.prev lbl with
           | [ pre ] when (Pcfg.branching pre) ->
-            let seq = Sigs.{ pre ; post = lbl } in
+            let seq = Memory.{ pre ; post = lbl } in
             let upd = lang#updates seq in
             Some(pre,upd)
           | _ -> None in
         let after = match Pcfg.next lbl with
           | [ post ] ->
-            let seq = Sigs.{ pre = lbl ; post } in
+            let seq = Memory.{ pre = lbl ; post } in
             let upd = lang#updates seq in
             Some(lbl,upd)
           | _ -> None in
diff --git a/src/plugins/wp/Pcond.mli b/src/plugins/wp/Pcond.mli
index eecd058b260cfd11fb639c3eff2e26405311f1d7..87746c20afa63566a14e8d315ab7ef473f02ecf9 100644
--- a/src/plugins/wp/Pcond.mli
+++ b/src/plugins/wp/Pcond.mli
@@ -123,9 +123,9 @@ class state :
 
     method domain : Vars.t
     method label_at : id:int -> Pcfg.label
-    method updates : Pcfg.label Sigs.sequence -> Sigs.update Bag.t
+    method updates : Pcfg.label Memory.sequence -> Memory.update Bag.t
     method pp_at : Format.formatter -> Pcfg.label -> unit
-    method pp_update : Pcfg.label -> Format.formatter -> Sigs.update -> unit
+    method pp_update : Pcfg.label -> Format.formatter -> Memory.update -> unit
     method pp_value : Format.formatter -> term -> unit
   end
 
diff --git a/src/plugins/wp/RegionAnalysis.ml b/src/plugins/wp/RegionAnalysis.ml
index f78d6e6eb2d71d3ccb8b6dc0f0aef72b193b5d94..f08f6c74a382e6bdb16c42d5c87b107eda6a7411 100644
--- a/src/plugins/wp/RegionAnalysis.ml
+++ b/src/plugins/wp/RegionAnalysis.ml
@@ -36,18 +36,18 @@ let id region = Region.uid (get_map ()) region
 let of_id id =
   try Some (Region.find (get_map ()) id)
   with Not_found -> None
+let compare r1 r2 = Int.compare (Region.id r1) (Region.id r2)
+let pretty fmt r = Format.fprintf fmt "R%03d" @@ Region.id r
 
-module Type =
+module R =
 struct
   type t = region
-  let hash = Region.id
-  let equal r1 r2 = (Region.id r1 = Region.id r2)
-  let compare r1 r2 = Int.compare (Region.id r1) (Region.id r2)
-  let pretty fmt r = Format.fprintf fmt "R%03d" @@ Region.id r
+  let compare = compare
+  let pretty = pretty
 end
 
 (* Keeping track of the decision to apply which memory model to each region *)
-module Kind = WpContext.Generator(Type)
+module Kind = WpContext.Generator(R)
     (struct
       open MemRegion
       let name = "Wp.RegionAnalysis.Kind"
@@ -68,7 +68,7 @@ module Kind = WpContext.Generator(Type)
         | None -> Garbled
     end)
 
-module Name = WpContext.Generator(Type)
+module Name = WpContext.Generator(R)
     (struct
       let name = "Wp.RegionAnalysis.Name"
       type key = region
diff --git a/src/plugins/wp/Sigma.ml b/src/plugins/wp/Sigma.ml
index ef6b8259547515f5ccbe802adefe1227e0d21309..fde69a2e0629a902264b2783faa10db3bcb8c149 100644
--- a/src/plugins/wp/Sigma.ml
+++ b/src/plugins/wp/Sigma.ml
@@ -20,168 +20,286 @@
 (*                                                                        *)
 (**************************************************************************)
 
+module F = Lang.F
+
+(* -------------------------------------------------------------------------- *)
+(* --- Generic Chunk Type                                                 --- *)
+(* -------------------------------------------------------------------------- *)
+
+module type ChunkType =
+sig
+  type t
+  val self : string
+  val hash : t -> int
+  val equal : t -> t -> bool
+  val compare : t -> t -> int
+  val pretty : Format.formatter -> t -> unit
+  val tau_of_chunk : t -> F.tau
+  val basename_of_chunk : t -> string
+  val is_init : t -> bool
+  val is_primary : t -> bool
+  val is_framed : t -> bool
+end
+
 (* -------------------------------------------------------------------------- *)
 (* --- Generic Sigma Factory                                              --- *)
 (* -------------------------------------------------------------------------- *)
 
-open Lang.F
+type ckind = ..
 
-module Make
-    (C : Sigs.Chunk)
-    (H : Qed.Collection.S with type t = C.t) :
-  Sigs.Sigma with type chunk = C.t
-              and module Chunk = H =
+module type Tag =
+sig
+  val tag : int
+  include ChunkType with type t := ckind
+end
+
+(* memory chunks *)
+type chunk = { tag : (module Tag) ; ckind : ckind }
+let ckind c = c.ckind
+
+module Chunk : ChunkType with type t = chunk =
 struct
+  type t = chunk
+  let self = "Core"
+  let hash c =
+    let module T = (val c.tag) in Qed.Hcons.hash_pair T.tag @@ T.hash c.ckind
+  let equal a b =
+    let module A = (val a.tag) in
+    let module B = (val b.tag) in
+    A.tag = B.tag && A.equal a.ckind b.ckind
+  let compare a b =
+    let module A = (val a.tag) in
+    let module B = (val b.tag) in
+    match A.is_primary a.ckind, B.is_primary b.ckind with
+    | true, false -> (-1)
+    | false, true -> (+1)
+    | true, true | false, false ->
+      let cmp = Int.compare A.tag B.tag in
+      if cmp <> 0 then cmp else A.compare a.ckind b.ckind
 
-  type chunk = C.t
-  module Chunk = H
-  type domain = H.Set.t
-
-  let empty = H.Set.empty
-  let union = H.Set.union
-
-  type t = { id : int ; mutable map : var H.map }
-
-  let id = ref 0 (* for debugging purpose *)
-  let build map = let k = !id in incr id ; { id = k ; map = map }
-  let create () = build H.Map.empty
-  let copy s = build s.map
-
-  let newchunk c =
-    Lang.freshvar ~basename:(C.basename_of_chunk c) (C.tau_of_chunk c)
-
-  let merge a b =
-    let pa = ref Passive.empty in
-    let pb = ref Passive.empty in
-    let merge_chunk c x y =
-      if Var.equal x y then x else
-        let z = newchunk c in
-        pa := Passive.bind ~fresh:z ~bound:x !pa ;
-        pb := Passive.bind ~fresh:z ~bound:y !pb ;
-        z in
-    let w = H.Map.union merge_chunk a.map b.map in
-    build w , !pa , !pb
-
-  type kind =
-    | Used of Lang.F.var
-    | Unused
-
-  let merge_list l =
-    (* Get a map of the chunks (the data is not important) *)
-    let union = List.fold_left (fun acc e -> H.Map.union (fun _ v1 _ -> v1) acc e.map) H.Map.empty l in
-    (* The goal is to build a matrix chunk -> elt of the list -> Used/Unused
-    *)
-    (* Set the data of the map to []. *)
-    let union = H.Map.map (fun _ -> []) union in
-    (* For each elements of the list tell if each chunk is used *)
-    let merge _ m e =
-      match m, e with
-      | Some m, Some e -> Some (Used e::m)
-      | Some m, None -> Some (Unused::m)
-      | None, _ -> assert false in
-    let union = List.fold_left (fun acc e -> H.Map.merge merge acc e.map) union
-        (* important so that the list in the map are in the correct order *)
-        (List.rev l) in
-    (* Build the passive for each element of the list, and the final domain *)
-    let p = ref (List.map (fun _ -> Passive.empty) l) in
-    let map c l =
-      match List.filter (fun x -> not (Unused = x)) l with
-      | [] -> assert false
-      (* If all the sigmas use the same variable *)
-      | (Used a)::l when List.for_all (function | Unused -> true | Used x -> Var.equal x a) l ->
-        a
-      | _ ->
-        let z = newchunk c in
-        let map2 p = function
-          | Unused -> p
-          | Used a -> Passive.bind ~fresh:z ~bound:a p
-        in
-        p := List.map2 map2 !p l;
-        z
-    in
-    let union = H.Map.mapi map union in
-    build union , !p
-
-  let choose a b =
-    let merge_chunck _ x y = if Var.compare x y <= 0 then x else y in
-    build (H.Map.union merge_chunck a.map b.map)
-
-  let get w c =
-    try H.Map.find c w.map
-    with Not_found ->
-      let x = newchunk c in
-      w.map <- H.Map.add c x w.map ; x
-
-  let mem w c = H.Map.mem c w.map
-
-  let join a b =
-    if a == b then Passive.empty else
-      let p = ref Passive.empty in
-      H.Map.iter2
-        (fun chunk x y ->
-           match x,y with
-           | Some x , Some y -> p := Passive.join x y !p
-           | Some x , None -> b.map <- H.Map.add chunk x b.map
-           | None , Some y -> a.map <- H.Map.add chunk y a.map
-           | None , None -> ())
-        a.map b.map ; !p
-
-  let assigned ~pre ~post written =
-    let p = ref Bag.empty in
-    H.Map.iter2
-      (fun chunk x y ->
-         if not (H.Set.mem chunk written) then
-           match x,y with
-           | Some x , Some y when x != y ->
-             p := Bag.add (p_equal (e_var x) (e_var y)) !p
-           | Some x , None -> post.map <- H.Map.add chunk x post.map
-           | None , Some y -> pre.map <- H.Map.add chunk y pre.map
-           | _ -> ())
-      pre.map post.map ; !p
+  let pretty fmt c =
+    let module T = (val c.tag) in T.pretty fmt c.ckind
+  let tau_of_chunk c =
+    let module T = (val c.tag) in T.tau_of_chunk c.ckind
+  let basename_of_chunk c =
+    let module T = (val c.tag) in T.basename_of_chunk c.ckind
+  let is_init c =
+    let module T = (val c.tag) in T.is_init c.ckind
+  let is_primary c =
+    let module T = (val c.tag) in T.is_primary c.ckind
+  let is_framed c =
+    let module T = (val c.tag) in T.is_framed c.ckind
+end
+
+module Heap = Qed.Collection.Make(Chunk)
+module Domain = Heap.Set
 
-  let value w c = e_var (get w c)
+(* -------------------------------------------------------------------------- *)
+(* --- Domain                                                             --- *)
+(* -------------------------------------------------------------------------- *)
 
-  let iter f w = H.Map.iter f w.map
-  let iter2 f w1 w2 = H.Map.iter2 f w1.map w2.map
+type domain = Domain.t
+let empty = Domain.empty
+let union = Domain.union
 
-  let havoc w xs =
-    let ys = H.Set.mapping newchunk xs in
-    build (H.Map.union (fun _c _old y -> y) w.map ys)
+(* -------------------------------------------------------------------------- *)
+(* --- State                                                              --- *)
+(* -------------------------------------------------------------------------- *)
 
-  let havoc_chunk w c =
+type state = chunk F.Tmap.t
+
+let apply f (s:state) : state =
+  F.Tmap.fold (fun t c s -> F.Tmap.add (f t) c s) s F.Tmap.empty
+
+(* -------------------------------------------------------------------------- *)
+(* --- Sigma                                                              --- *)
+(* -------------------------------------------------------------------------- *)
+
+type sigma = F.var Heap.Map.t ref
+
+let create () = ref Heap.Map.empty
+let copy sigma = ref !sigma
+let newchunk c =
+  let module T = (val c.tag) in
+  let basename = T.basename_of_chunk c.ckind in
+  let tau = T.tau_of_chunk c.ckind in
+  Lang.freshvar ~basename tau
+
+let mem (sigma : sigma) c = Heap.Map.mem c !sigma
+let domain (sigma : sigma) : domain = Heap.Map.domain !sigma
+let state (sigma : sigma) : state =
+  let s = ref F.Tmap.empty in
+  Heap.Map.iter (fun c x -> s := F.Tmap.add (F.e_var x) c !s) !sigma ; !s
+
+let get sigma c =
+  let s = !sigma in
+  try Heap.Map.find c s
+  with Not_found ->
     let x = newchunk c in
-    build (H.Map.add c x w.map)
-
-  let havoc_any ~call w =
-    let framer c x = if call && C.is_framed c then x else newchunk c in
-    build (H.Map.mapi framer w.map)
-
-  let remove_chunks w xs =
-    build (H.Map.filter (fun c _ -> not (H.Set.mem c xs)) w.map)
-
-  let domain w = H.Map.domain w.map
-
-  let pretty fmt w =
-    begin
-      Format.fprintf fmt "@[<hov 2>@@%s%d[" C.self w.id ;
-      H.Map.iter
-        (fun c x -> Format.fprintf fmt "@ %a:%a" C.pretty c Var.pretty x) w.map ;
-      Format.fprintf fmt " ]@]" ;
-    end
-
-  let writes seq =
-    let writes = ref Chunk.Set.empty in
-    iter2
-      (fun chunk u v ->
-         let written =
-           match u,v with
-           | Some x , Some y -> not (Var.equal x y)
-           | None , Some _ -> true
-           | Some _ , None -> false (* no need to create a new so it is the same *)
-           | None, None -> assert false
-         in
-         if written then writes := Chunk.Set.add chunk !writes
-      ) seq.Sigs.pre seq.Sigs.post ;
-    !writes
+    sigma := Heap.Map.add c x s ; x
+
+let value sigma c = F.e_var @@ get sigma c
 
+let merge a b =
+  let pa = ref Passive.empty in
+  let pb = ref Passive.empty in
+  let merge_chunk c x y =
+    if F.Var.equal x y then x else
+      let z = newchunk c in
+      pa := Passive.bind ~fresh:z ~bound:x !pa ;
+      pb := Passive.bind ~fresh:z ~bound:y !pb ;
+      z in
+  let merged = Heap.Map.union merge_chunk !a !b in
+  ref merged , !pa, !pb
+
+let choose a b =
+  let merge_chunck _ x y = if F.Var.compare x y <= 0 then x else y in
+  ref @@ Heap.Map.union merge_chunck !a !b
+
+let join a b =
+  if a == b then Passive.empty else
+    let p = ref Passive.empty in
+    Heap.Map.iter2
+      (fun chunk x y ->
+         match x,y with
+         | Some x , Some y -> p := Passive.join x y !p
+         | Some x , None -> b := Heap.Map.add chunk x !b
+         | None , Some y -> a := Heap.Map.add chunk y !a
+         | None , None -> ())
+      !a !b ; !p
+
+let havoc sigma domain =
+  let ys = Domain.mapping newchunk domain in
+  ref @@ Heap.Map.union (fun _v _old y -> y) !sigma ys
+
+let remove_chunks sigma domain =
+  ref @@ Heap.Map.filter (fun c _ -> not @@ Domain.mem c domain) !sigma
+
+let havoc_chunk sigma c =
+  let x = newchunk c in
+  ref @@ Heap.Map.add c x !sigma
+
+let is_init c =
+  let module T = (val c.tag) in T.is_init c.ckind
+
+let is_framed c =
+  let module T = (val c.tag) in T.is_framed c.ckind
+
+let havoc_any ~call sigma =
+  let framer c x = if call && is_framed c then x else newchunk c in
+  ref @@ Heap.Map.mapi framer !sigma
+
+(* Merge All *)
+
+type usage = Used of F.var | Unused
+
+let merge_list (ls : sigma list) : sigma * Passive.t list =
+  (* Get a map of the chunks (the data is not important) *)
+  let domain =
+    Heap.Map.map (fun _ -> []) @@
+    let dunion _ c _ = c in
+    List.fold_left
+      (fun acc s -> Heap.Map.merge dunion acc !s)
+      Heap.Map.empty ls in
+  (* Accumulate usage for each c, preserving the order of the list *)
+  let usage =
+    let umerge _ c e =
+      match c, e with
+      | Some c, Some x -> Some (Used x :: c)
+      | Some c, None -> Some (Unused :: c)
+      | None, _ -> assert false in
+    List.fold_left
+      (fun acc s -> Heap.Map.merge umerge acc !s)
+      domain (List.rev ls) in
+  (* Build the passive for each sigma of the list, and the final sigma *)
+  let passives = ref @@ List.map (fun _ -> Passive.empty) ls in
+  let choose c uses =
+    let used = function Used _ -> true | Unused -> false in
+    let compatible x = function Used y -> F.Var.equal x y | Unused -> true in
+    match List.filter used uses with
+    | [] -> assert false
+    | Used x :: others when List.for_all (compatible x) others -> x
+    | _ ->
+      let fresh = newchunk c in
+      let join pa = function
+        | Unused -> pa
+        | Used bound -> Passive.bind ~fresh ~bound pa in
+      passives := List.map2 join !passives uses ; fresh
+  in ref @@ Heap.Map.mapi choose usage, !passives
+
+(* Effects *)
+
+let writes ~(pre:sigma) ~(post:sigma) : domain =
+  let domain = ref Domain.empty in
+  Heap.Map.iter2
+    (fun c u v ->
+       let written =
+         match u,v with
+         | Some x , Some y -> not @@ F.Var.equal x y
+         | None , Some _ -> true
+         | Some _ , None -> false (* no need to create a new *)
+         | None, None -> assert false
+       in if written then domain := Domain.add c !domain
+    ) !pre !post ;
+  !domain
+
+let assigned ~(pre:sigma) ~(post:sigma) written =
+  let p = ref Bag.empty in
+  Heap.Map.iter2
+    (fun chunk x y ->
+       if not (Domain.mem chunk written) then
+         match x,y with
+         | Some x , Some y when x != y ->
+           p := Bag.add (F.p_equal (F.e_var x) (F.e_var y)) !p
+         | Some x , None -> post := Heap.Map.add chunk x !post
+         | None , Some y -> pre := Heap.Map.add chunk y !pre
+         | _ -> ())
+    !pre !post ; !p
+
+let iter f s = Heap.Map.iter f !s
+let iter2 f a b = Heap.Map.iter2 f !a !b
+
+let pretty fmt sigma =
+  begin
+    Format.fprintf fmt "{@[<hov 2>" ;
+    Heap.Map.iter
+      (fun c x ->
+         let module T = (val c.tag) in
+         Format.fprintf fmt "@ %a:%a" T.pretty c.ckind F.Var.pretty x)
+      !sigma ;
+    Format.fprintf fmt "@]}" ;
+  end
+
+(* -------------------------------------------------------------------------- *)
+(* --- Chunks                                                             --- *)
+(* -------------------------------------------------------------------------- *)
+
+module Make(C : ChunkType) =
+struct
+  type ckind += Mu of C.t
+  module T =
+  struct
+    let self = C.self
+    let tag = Obj.Extension_constructor.id [%extension_constructor Mu]
+    let map f = function Mu m -> f m | _ -> assert false
+    let map2 f a b =
+      match a,b with Mu a, Mu b -> f a b | _ -> assert false
+    let hash = map C.hash
+    let equal = map2 C.equal
+    let compare = map2 C.compare
+    let is_init = map C.is_init
+    let is_primary = map C.is_primary
+    let is_framed = map C.is_framed
+    let tau_of_chunk = map C.tau_of_chunk
+    let basename_of_chunk = map C.basename_of_chunk
+    let pretty fmt = map (C.pretty fmt)
+  end
+  let tag = (module T : Tag)
+  let chunk c = { tag ; ckind = Mu c }
+  let mem sigma c = mem sigma @@ chunk c
+  let get sigma c = get sigma @@ chunk c
+  let value sigma c = value sigma @@ chunk c
+  let singleton c = Domain.singleton @@ chunk c
 end
+
+(* -------------------------------------------------------------------------- *)
diff --git a/src/plugins/wp/Sigma.mli b/src/plugins/wp/Sigma.mli
index 6d1e8fef72c126c6c64ab09b372830e808222f47..94f5cc611dc45c0a44ba769b59655776cf70d819 100644
--- a/src/plugins/wp/Sigma.mli
+++ b/src/plugins/wp/Sigma.mli
@@ -20,12 +20,225 @@
 (*                                                                        *)
 (**************************************************************************)
 
+module F = Lang.F
+
 (* -------------------------------------------------------------------------- *)
-(* --- Generic Sigma Factory                                              --- *)
+(* --- Memory Environments                                                --- *)
 (* -------------------------------------------------------------------------- *)
 
-module Make
-    (C : Sigs.Chunk)
-    (H : Qed.Collection.S with type t = C.t) :
-  Sigs.Sigma with type chunk = C.t
-              and module Chunk = H
+(** Memory Environments.
+
+    The concrete memory is partionned into a vector of abstract data.
+    Each component of the partition is called a {i memory chunk} and
+    holds an abstract representation of some part of the memory.
+
+    A sigma assigns a logical variable to each {i memory chunk}, assigned
+    lazily, that represents the contents of this slice of memory at a given
+    program point.
+
+    Remark: memory chunks are not required to be independant from each other,
+    provided the memory model implementation is consistent with the chosen
+    representation. Conversely, a given object might be represented by
+    several memory chunks.
+*)
+
+(** Memory chunk types.
+
+    The type of chunks is extensible.
+    Each memory model extends this type with of its {i own} memory chunks
+    by functor {!Make}. *)
+type chunk
+
+(** Memory chunks at a given program point.
+
+    This is a collection of variables indexed by chunks.
+    New chunks are generated from the context pool of {!Lang.freshvar}. *)
+type sigma
+
+(** {1 Chunks API} *)
+
+module type ChunkType =
+sig
+
+  type t
+
+  val self : string
+  (** Chunk type name for datatype registration. *)
+
+  val hash : t -> int
+  val equal : t -> t -> bool
+  val compare : t -> t -> int
+  val pretty : Format.formatter -> t -> unit
+
+  val tau_of_chunk : t -> F.tau
+  (** The type of data hold in a chunk. *)
+
+  val basename_of_chunk : t -> string
+  (** Used when generating fresh variables for a chunk. *)
+
+  val is_init : t -> bool
+  (** Whether the chunk tracks memory initialization (used for filtering). *)
+
+  val is_primary : t -> bool
+  (** Used to sort memory chunk parameters in ACSL symbols.
+
+      Single chunks represents values of variables or individual l-values.
+      They are sorted last among memory chunk parameters. *)
+
+  val is_framed : t -> bool
+  (** Whether the chunk is local to a function call.
+
+      Means the chunk is separated from anyother call side-effects.
+      If [true], entails that a function assigning everything can not modify
+      the chunk. Only used for optimisation, it would be safe to always
+      return [false]. *)
+end
+
+(** Uniform access to chunks features. *)
+module Chunk : ChunkType with type t = chunk
+
+(** Uniform Sets and Maps of chunks. *)
+module Heap : Qed.Collection.S with type t = chunk
+module Domain = Heap.Set
+
+(** Memory footprint, ie. set of memory model chunk types. *)
+type domain = Domain.t
+
+(** Internal representation of chunks (chunk-kind).
+    This type can only be extended via functor {!Make}. *)
+type ckind = private ..
+
+(** Access to internal chunk representation. *)
+val ckind : chunk -> ckind
+
+(** Chunk Extension functor.
+
+    This functor promote an individual chunk type of
+    some memory model into an uniform chunk.
+*)
+module Make(C : ChunkType) :
+sig
+  (** Chunk Extension. *)
+  type ckind += Mu of C.t
+
+  val chunk : C.t -> chunk
+  (** Individual promotion of a model chunk to a uniform chunk. *)
+
+  val singleton : C.t -> domain
+  (** Promoted [Domain.singleton]. *)
+
+  val mem : sigma -> C.t -> bool
+  (** Specialized [Sigma.mem], equivalent to [Sigma.mem s (chunk c)]. *)
+
+  val get : sigma -> C.t -> F.var
+  (** Specialized [Sigma.get], equivalent to [Sigma.get s (chunk c)]. *)
+
+  val value : sigma -> C.t -> F.term
+  (** Specialized [Sigma.value], equivalent to [Sigma.value s (chunk c)]. *)
+
+end
+
+(** {1 Sigma API} *)
+
+val create : unit -> sigma
+(** Initially empty environment. *)
+
+val pretty : Format.formatter -> sigma -> unit
+(** For debugging purpose *)
+
+val mem : sigma -> chunk -> bool
+(** Whether a chunk has been assigned. *)
+
+val get : sigma -> chunk -> F.var
+(** Lazily get the variable for a chunk. *)
+
+val value : sigma -> chunk -> F.term
+(** Same as [Lang.F.e_var] of [get]. *)
+
+val copy : sigma -> sigma
+(** Duplicate the environment. Fresh chunks in the copy
+    are {i not} duplicated into the source environment. *)
+
+val join : sigma -> sigma -> Passive.t
+(** Make two environment pairwise equal {i via} the passive form.
+
+    Missing chunks in one environment are added with the corresponding
+    variable of the other environment. When both environments don't agree
+    on a chunk, their variables are added to the passive form. *)
+
+val assigned : pre:sigma -> post:sigma -> domain -> F.pred Bag.t
+(** Make chunks equal outside of some domain.
+
+    This is similar to [join], but outside the given footprint of an
+    assigns clause. Although, the function returns the equality
+    predicates instead of a passive form.
+
+    Like in [join], missing chunks are reported from one side to the
+    other one, and common chunks are added to the equality bag. *)
+
+val choose : sigma -> sigma -> sigma
+(** Make the union of each sigma, choosing the minimal variable
+    in case of conflict.
+    Both initial environments are kept unchanged. *)
+
+val merge : sigma -> sigma -> sigma * Passive.t * Passive.t
+(** Make the union of each sigma, choosing a {i new} variable for
+    each conflict, and returns the corresponding joins.
+    Both initial environments are kept unchanged. *)
+
+val merge_list : sigma list -> sigma * Passive.t list
+(** Same than {!merge} but for a list of sigmas. Much more efficient
+    than folding merge step by step. *)
+
+val iter : (chunk -> F.var -> unit) -> sigma -> unit
+(** Iterates over the chunks and associated variables already
+    accessed so far in the environment. *)
+
+val iter2 :
+  (chunk -> F.var option -> F.var option -> unit) ->
+  sigma -> sigma -> unit
+(** Same as [iter] for both environments. *)
+
+val havoc_chunk : sigma -> chunk -> sigma
+(** Generate a new fresh variable for the given chunk. *)
+
+val havoc : sigma -> domain -> sigma
+(** All the chunks in the provided footprint are generated and made fresh.
+
+    Existing chunk variables {i outside} the footprint are copied into the new
+    environment. The original environement itself is kept unchanged. More
+    efficient than iterating [havoc_chunk] over the footprint.
+*)
+
+val havoc_any : call:bool -> sigma -> sigma
+(** All the chunks are made fresh. As an optimisation,
+    when [~call:true] is set, only non-local chunks are made fresh.
+    Local chunks are those for which [Chunk.is_frame] returns [true]. *)
+
+val remove_chunks : sigma -> domain -> sigma
+(** Return a copy of the environment where chunks in the footprint
+    have been removed. Keep the original environment unchanged. *)
+
+val is_init : chunk -> bool
+(** Shortcut to [Chunk.is_init] *)
+
+val is_framed : chunk -> bool
+(** Shortcut to [Chunk.is_framed] *)
+
+val domain : sigma -> domain
+(** Footprint of a memory environment.
+    That is, the set of accessed chunks so far in the environment. *)
+
+val union : domain -> domain -> domain
+(** Same as [Domain.union] *)
+
+val empty : domain
+(** Same as [Domain.empty] *)
+
+val writes : pre:sigma -> post:sigma -> domain
+(** [writes s] indicates which chunks are new in [s.post] compared
+    to [s.pre]. *)
+
+type state = chunk F.Tmap.t
+val state : sigma -> state
+val apply : (F.term -> F.term) -> state -> state
diff --git a/src/plugins/wp/StmtSemantics.ml b/src/plugins/wp/StmtSemantics.ml
index d5a521e7e185cc83314a30b86e30e1b3bd8bbd2b..fa19b5e152d237a12ce1552651de4fb543bd2580 100644
--- a/src/plugins/wp/StmtSemantics.ml
+++ b/src/plugins/wp/StmtSemantics.ml
@@ -20,20 +20,19 @@
 (*                                                                        *)
 (**************************************************************************)
 
-open Sigs
+open Memory
 open Cil_types
 open Cil_datatype
 open Clabels
+module Cfg = CfgCompiler
 
 let not_yet = Wp_parameters.not_yet_implemented
 
-module Make(Compiler:Sigs.Compiler) =
+module Make(Compiler:Memory.Compiler) =
 struct
 
   module Compiler = Compiler
-  module Cfg = CfgCompiler.Cfg(Compiler.M.Sigma)
   module M = Compiler.M
-  module Sigma = Compiler.M.Sigma
   module C  = Compiler.C
   module L = Compiler.L
   module A = Compiler.A
@@ -242,7 +241,7 @@ struct
   let mk_frame ~descr env =
     let nsigmas =
       LabelMap.fold
-        (fun _ (n : node) (nmap : M.sigma Cfg.Node.Map.t) ->
+        (fun _ (n : node) (nmap : Sigma.sigma Cfg.Node.Map.t) ->
            if Cfg.Node.Map.mem n nmap then nmap
            else Cfg.Node.Map.add n (Sigma.create ()) nmap)
         env.flow
@@ -271,14 +270,14 @@ struct
     let frame =
       L.mk_frame
         ~labels:lsigmas ~kf:env.kf
-        ~result:(Sigs.R_var env.result)
+        ~result:(Memory.R_var env.result)
         ~status:env.status
         ~formals ~descr ()
     in
     frame, nsigmas, lsigmas
 
   let pred
-    : env -> Sigs.polarity -> predicate -> _
+    : env -> Memory.polarity -> predicate -> _
     = fun env polarity p ->
       (* Format.printf "env.flow: %a@." *)
       (*   (Pretty_utils.pp_iter2 LabelMap.iter Label.pretty Cfg.Node.pp) *)
@@ -290,7 +289,7 @@ struct
         let pred = L.in_frame frame (L.pred polarity lenv) p in
         (* Remove the sigmas not used for the compilation, but here must stay *)
         let nsigmas = Cfg.Node.Map.filter (fun _ s ->
-            s == here || not (Sigma.Chunk.Set.is_empty (Sigma.domain s))
+            s == here || not (Sigma.Domain.is_empty (Sigma.domain s))
           ) nsigmas
         in
         (Cfg.P.create nsigmas pred)
@@ -311,7 +310,7 @@ struct
 
 
   let assume_
-    : env -> Sigs.polarity -> predicate -> paths
+    : env -> Memory.polarity -> predicate -> paths
     = fun env polarity p ->
       assume (pred env polarity p)
 
@@ -365,7 +364,7 @@ struct
 
       let exit_status (env:env) =
         let p = Lang.F.p_equal (Lang.F.e_var old_status) (Lang.F.e_var env.status) in
-        let s = M.Sigma.create () in
+        let s = Sigma.create () in
         let e = Cfg.E.create {pre=s;post=s} p in
         memory_effect (env @: Clabels.here) e (env @: Clabels.next)
       in
@@ -484,7 +483,7 @@ struct
     | None -> goto (env @: Clabels.here) (env @: Clabels.next)
     | Some region ->
       let domain = A.domain region in
-      let next = M.Sigma.havoc here domain in
+      let next = Sigma.havoc here domain in
       let seq = { pre = here; post = next } in
       let preds = A.apply_assigns seq region in
       memory_effect (env @: Clabels.here) (Cfg.E.create seq (Lang.F.p_conj preds)) (env @: Clabels.next)
@@ -525,8 +524,8 @@ struct
       match tr with
       | Skip | Enter { blocals = [] } | Leave { blocals = [] } ->
         goto (env @: Clabels.here) (env @: Clabels.next)
-      | Enter {blocals} -> scope env Sigs.Enter blocals
-      | Leave {blocals} -> scope env Sigs.Leave blocals
+      | Enter {blocals} -> scope env Memory.Enter blocals
+      | Leave {blocals} -> scope env Memory.Leave blocals
       | Return (r,_) -> return env r
       | Prop ({kind = Assert|Invariant} as a, _) ->
         let env = Logic_label.Map.fold
diff --git a/src/plugins/wp/StmtSemantics.mli b/src/plugins/wp/StmtSemantics.mli
index 36ce4a01a06047da4679ff79bcfd057c24c1a907..72ec68d3a9729d075c9f72d23c078b82eeb18508 100644
--- a/src/plugins/wp/StmtSemantics.mli
+++ b/src/plugins/wp/StmtSemantics.mli
@@ -22,102 +22,101 @@
 
 open Cil_types
 open Clabels
+module Cfg = CfgCompiler
 
-module Make(Compiler : Sigs.Compiler) :
-sig
+module Make : Memory.Compiler ->
+  sig
 
-  module Cfg : CfgCompiler.Cfg with module S = Compiler.M.Sigma
+    type node = Cfg.node
+    type goal = {
+      goal_pred : Cfg.P.t;
+      goal_prop : WpPropId.prop_id;
+    }
+    type cfg = Cfg.cfg
+    type paths = {
+      paths_cfg : cfg;
+      paths_goals : goal Bag.t;
+    }
 
-  type node = Cfg.node
-  type goal = {
-    goal_pred : Cfg.P.t;
-    goal_prop : WpPropId.prop_id;
-  }
-  type cfg = Cfg.cfg
-  type paths = {
-    paths_cfg : cfg;
-    paths_goals : goal Bag.t;
-  }
+    val goals_nodes: goal Bag.t -> Cfg.Node.Set.t
 
-  val goals_nodes: goal Bag.t -> Cfg.Node.Set.t
+    exception LabelNotFound of c_label
 
-  exception LabelNotFound of c_label
+    (** Compilation environment *)
 
-  (** Compilation environment *)
+    type env
 
-  type env
+    val empty_env : Kernel_function.t -> env
+    val bind : c_label -> node -> env -> env
 
-  val empty_env : Kernel_function.t -> env
-  val bind : c_label -> node -> env -> env
+    val result : env -> Lang.F.var
 
-  val result : env -> Lang.F.var
+    val (@^) : paths -> paths -> paths
+    (** Same as [Cfg.concat] *)
 
-  val (@^) : paths -> paths -> paths
-  (** Same as [Cfg.concat] *)
+    val (@*) : env -> ( c_label * node ) list -> env
+    (** fold bind *)
 
-  val (@*) : env -> ( c_label * node ) list -> env
-  (** fold bind *)
+    val (@:) : env -> c_label -> node
+    (** LabelMap.find with refined excpetion.
+        @raise LabelNotFound instead of [Not_found] *)
 
-  val (@:) : env -> c_label -> node
-  (** LabelMap.find with refined excpetion.
-      @raise LabelNotFound instead of [Not_found] *)
+    val (@-) : env -> (c_label -> bool) -> env
 
-  val (@-) : env -> (c_label -> bool) -> env
+    val sequence : (env -> 'a -> paths) -> env -> 'a list -> paths
+    (** Chain compiler by introducing fresh nodes between each element
+        of the list. For each consecutive [x;y] elements, a fresh node [n]
+        is created, and [x] is compiled with [Next:n] and [y] is compiled with
+        [Here:n]. *)
 
-  val sequence : (env -> 'a -> paths) -> env -> 'a list -> paths
-  (** Chain compiler by introducing fresh nodes between each element
-      of the list. For each consecutive [x;y] elements, a fresh node [n]
-      is created, and [x] is compiled with [Next:n] and [y] is compiled with
-      [Here:n]. *)
+    val choice : ?pre:c_label -> ?post:c_label ->
+      (env -> 'a -> paths) -> env -> 'a list -> paths
+    (** Chain compiler in parallel, between labels [~pre] and [~post], which
+        defaults to resp. [here] and [next].
+        The list of eventualities is exhastive, hence an [either] assumption
+        is also inserted. *)
 
-  val choice : ?pre:c_label -> ?post:c_label ->
-    (env -> 'a -> paths) -> env -> 'a list -> paths
-  (** Chain compiler in parallel, between labels [~pre] and [~post], which
-      defaults to resp. [here] and [next].
-      The list of eventualities is exhastive, hence an [either] assumption
-      is also inserted. *)
+    val parallel : ?pre:c_label -> ?post:c_label ->
+      (env -> 'a -> Cfg.C.t * paths) -> env -> 'a list -> paths
+    (** Chain compiler in parallel, between labels [~pre] and [~post], which
+        defaults to resp. [here] and [next].
+        The list of eventualities is exhastive, hence an [either] assumption
+        is also inserted. *)
 
-  val parallel : ?pre:c_label -> ?post:c_label ->
-    (env -> 'a -> Cfg.C.t * paths) -> env -> 'a list -> paths
-  (** Chain compiler in parallel, between labels [~pre] and [~post], which
-      defaults to resp. [here] and [next].
-      The list of eventualities is exhastive, hence an [either] assumption
-      is also inserted. *)
+    (** {2 Instructions Compilation}
 
-  (** {2 Instructions Compilation}
+        Each instruction or statement is typically compiled between
+        [Here] and [Next] nodes in the [flow]. [Pre], [Post] and [Exit] are
+        reserved for the entry and exit points of current function.
+        in [flow] are used when needed such as [Break] and [Continue] and
+        should be added before calling.
+    *)
 
-      Each instruction or statement is typically compiled between
-      [Here] and [Next] nodes in the [flow]. [Pre], [Post] and [Exit] are
-      reserved for the entry and exit points of current function.
-      in [flow] are used when needed such as [Break] and [Continue] and
-      should be added before calling.
-  *)
+    val set : env -> lval -> exp -> paths
+    val scope : env -> Memory.scope -> varinfo list -> paths
+    val instr : env -> instr -> paths
+    val return : env -> exp option -> paths
+    val assume : Cfg.P.t -> paths
 
-  val set : env -> lval -> exp -> paths
-  val scope : env -> Sigs.scope -> varinfo list -> paths
-  val instr : env -> instr -> paths
-  val return : env -> exp option -> paths
-  val assume : Cfg.P.t -> paths
+    val call_kf : env -> lval option -> kernel_function -> exp list -> paths
+    val call : env -> lval option -> exp -> exp list -> paths
 
-  val call_kf : env -> lval option -> kernel_function -> exp list -> paths
-  val call : env -> lval option -> exp -> exp list -> paths
+    (** {2 ACSL Compilation}  *)
 
-  (** {2 ACSL Compilation}  *)
+    val spec : env -> spec -> paths
+    val assume_ : env -> Memory.polarity -> predicate -> paths
+    val assigns : env -> assigns -> paths
+    val froms : env -> from list -> paths
 
-  val spec : env -> spec -> paths
-  val assume_ : env -> Sigs.polarity -> predicate -> paths
-  val assigns : env -> assigns -> paths
-  val froms : env -> from list -> paths
+    (** {2 Automata Compilation} *)
 
-  (** {2 Automata Compilation} *)
+    val automaton : env -> Interpreted_automata.automaton -> paths
 
-  val automaton : env -> Interpreted_automata.automaton -> paths
+    val init: is_pre_main:bool -> env -> paths
 
-  val init: is_pre_main:bool -> env -> paths
+    (** {2 Full Compilation}
 
-  (** {2 Full Compilation}
-
-      Returns the set of all paths for the function, with all proof
-      obligations. The returned node corresponds to the [Init] label. *)
-  val compute_kf: Kernel_function.t -> paths * node
-end
+        Returns the set of all paths for the function, with all proof
+        obligations. The returned node corresponds to the [Init] label. *)
+    val compute_kf: Kernel_function.t -> paths * node
+  end
diff --git a/src/plugins/wp/cfgWP.ml b/src/plugins/wp/cfgWP.ml
index 8f36f6cedd4c664e71085984991fe5f13297d593..65d2c9a6d3b79076006cdfa79e04771c2365a934 100644
--- a/src/plugins/wp/cfgWP.ml
+++ b/src/plugins/wp/cfgWP.ml
@@ -30,10 +30,10 @@ open Cil_datatype
 open WpPropId
 open Clabels
 open Qed
-open Lang
 open Lang.F
-open Sigs
 open Wpo
+open Memory
+open Sigma
 
 module type VCgen =
 sig
@@ -43,7 +43,7 @@ sig
   val compile_wp : Wpo.index -> t_prop -> Wpo.t Bag.t
 end
 
-module VC( C : Sigs.Compiler ) : VCgen =
+module VC( C : Memory.Compiler ) : VCgen =
 struct
 
   open C
@@ -51,8 +51,6 @@ struct
   module V = Vars
   module P = WpPropId.PropId
 
-  let state = Mstate.create (module M)
-
   type target =
     | Gprop of P.t
     | Geffect of P.t * Stmt.t * Mcfg.effect_source
@@ -140,7 +138,7 @@ struct
     e_pid : P.t ; (* Assign Property *)
     e_post : bool ; (* Requires post effects (loop-assigns or post-assigns) *)
     e_label : c_label ; (* scope for collection *)
-    e_valid : L.sigma ; (* sigma where locations are filtered for validity *)
+    e_valid : sigma ; (* sigma where locations are filtered for validity *)
     e_region : L.region ; (* expected from spec *)
     e_warn : Warning.Set.t ; (* from translation *)
   }
@@ -182,7 +180,7 @@ struct
   }
 
   type t_prop = {
-    sigma : L.sigma option ;
+    sigma : sigma option ;
     effects : Eset.t ;
     vcs : vc Splitter.t Gmap.t ;
   }
@@ -441,7 +439,7 @@ struct
     let frame = wenv.frame in
     let init = L.mem_at_frame frame Clabels.init in
     let domain = Sigma.domain init in
-    not (M.Heap.Set.is_empty domain)
+    not (Domain.is_empty domain)
 
   let merge wenv wp1 wp2 =
     L.in_frame wenv.frame
@@ -741,7 +739,7 @@ struct
         | None | Some { labels=[] } -> None
         | Some { labels = lbl::_ } ->
           Some (Pretty_utils.to_string Printer.pp_label lbl) in
-      let state = Mstate.state state sigma in
+      let state = Mstate.create (module M) sigma in
       gmap (state_vc ?descr ?stmt sigma state) vcs
     with Not_found -> vcs
 
@@ -831,8 +829,8 @@ struct
           | Warning.Result(r_warn,Some stored) ->
             (* R-Value and effects has been translated *)
             let warn = Warning.Set.union l_warn r_warn in
-            let ft = M.Heap.Set.fold_sorted
-                (fun chunk ft -> M.Sigma.get seq.post chunk :: ft) dom []
+            let ft = Domain.fold_sorted
+                (fun chunk ft -> Sigma.get seq.post chunk :: ft) dom []
             in
             let update vc =
               if List.exists (occurs_vc vc) ft
@@ -1285,7 +1283,7 @@ struct
       let tr = Cil.typeOfLval lv in
       let lr = C.lval dummy lv in
       Some (M.domain (Ctypes.object_of tr) lr)
-    | None -> Some (M.Heap.Set.empty)
+    | None -> Some Domain.empty
 
   let cc_call_domain env0 kf es = function
     | WritesAny -> None
diff --git a/src/plugins/wp/tests/wp/oracle_qualif/stmtcompiler_test_rela.res.oracle b/src/plugins/wp/tests/wp/oracle_qualif/stmtcompiler_test_rela.res.oracle
index df440421ab688b2c1e7168ea71419e664d401eb6..f7db5f42ae02f0d2f2c06889ef8883c463e41919 100644
--- a/src/plugins/wp/tests/wp/oracle_qualif/stmtcompiler_test_rela.res.oracle
+++ b/src/plugins/wp/tests/wp/oracle_qualif/stmtcompiler_test_rela.res.oracle
@@ -76,7 +76,7 @@ Sequent: Assume {
              Then { Have: (node_17=true). }
              Else { Have: (node_18=true). }
            }
-           (* Entering scope [tmp]: effect: (not Init_tmp_3) /\ (not ta_tmp_6) /\ ta_tmp_7 *)
+           (* Entering scope [tmp]: effect: (not Init_tmp_2) /\ (not ta_tmp_6) /\ ta_tmp_7 *)
            If (node_19=true)
            Then {
              Have: (Init_tmp_1=false) /\ (ta_tmp_4=true) /\ (ta_tmp_5=false).
diff --git a/src/plugins/wp/tests/wp/stmtcompiler_test.ml b/src/plugins/wp/tests/wp/stmtcompiler_test.ml
index 4ab9acd47ce62ac91057617cf0fe1578d02c0aa6..01a8adc35500b258f6fa27f6706aa789a41c7ff1 100644
--- a/src/plugins/wp/tests/wp/stmtcompiler_test.ml
+++ b/src/plugins/wp/tests/wp/stmtcompiler_test.ml
@@ -2,7 +2,7 @@
 
 open Wp
 open Factory
-open Sigs
+open Memory
 
 let mode = `Tree
 
@@ -95,7 +95,7 @@ let run () =
       begin fun () ->
         let automaton = Interpreted_automata.Compute.get_automaton ~annotations:true kf in
         (* Format.printf "@[%s body cil:%a@]@." fct Printer.pp_block block; *)
-        let seq = {Sigs.pre = Cfg.node (); post = Cfg.node ()} in
+        let seq = {Memory.pre = Cfg.node (); post = Cfg.node ()} in
         let env = Compiler.empty_env kf  in
         let env = Compiler.(env @* [Clabels.here,seq.pre;
                                     Clabels.next,seq.post]) in
diff --git a/src/plugins/wp/tests/wp/stmtcompiler_test_rela.ml b/src/plugins/wp/tests/wp/stmtcompiler_test_rela.ml
index 891afaeb731f32e9760851a8b335438cf2f17072..62b9efeec71b9f26d88d3dff40a904da6d6a0e0d 100644
--- a/src/plugins/wp/tests/wp/stmtcompiler_test_rela.ml
+++ b/src/plugins/wp/tests/wp/stmtcompiler_test_rela.ml
@@ -2,7 +2,7 @@
 
 open Wp
 open Factory
-open Sigs
+open Memory
 
 let run () =
   let setup : Factory.setup = { mheap = Hoare;
@@ -14,7 +14,7 @@ let run () =
   let model = Factory.instance setup driver in
   let module C = (val (Factory.compiler setup.mheap setup.mvar)) in
   let module Compiler = StmtSemantics.Make(C) in
-  let module Cfg = Compiler.Cfg in
+  let module Cfg = CfgCompiler in
 
   let provers =
     List.fold_right
@@ -101,27 +101,27 @@ let run () =
         let formal = List.hd (fct.sformals) in
 
         (*First call*)
-        let seq1 = {Sigs.pre = Cfg.node (); post = Cfg.node ()} in
+        let seq1 = {Memory.pre = Cfg.node (); post = Cfg.node ()} in
         let env1 = Compiler.empty_env kf  in
         let env1 = Compiler.(env1 @* [Clabels.here,seq1.pre;Clabels.next,seq1.post]) in
         let path1 = Compiler.automaton env1 block in
         let cfg1 = path1.Compiler.paths_cfg in
         let node1 = Cfg.T.init' seq1.pre (reads_formal formal) in
         let (_,sigma1,sequence1) =
-          Compiler.Cfg.compile seq1.pre
+          Cfg.compile seq1.pre
             (Cfg.Node.Set.singleton seq1.post) (Cfg.T.reads node1) cfg1 in
         let node1 = Cfg.T.relocate sigma1 node1 in
         let term_1 = Cfg.T.get node1 in
 
         (*Seconde call*)
-        let seq2 = {Sigs.pre = Cfg.node (); post = Cfg.node ()} in
+        let seq2 = {Memory.pre = Cfg.node (); post = Cfg.node ()} in
         let env2 = Compiler.empty_env kf  in
         let env2 = Compiler.(env2 @* [Clabels.here,seq2.pre;Clabels.next,seq2.post]) in
         let path2 = Compiler.automaton env2 block in
         let cfg2 = path2.Compiler.paths_cfg in
         let node2 = Cfg.T.init' seq2.pre (reads_formal formal) in
         let (_,sigma2,sequence2) =
-          Compiler.Cfg.compile
+          Cfg.compile
             seq2.pre (Cfg.Node.Set.singleton seq2.post)
             (Cfg.T.reads node2) cfg2 in
         let node2 = Cfg.T.relocate sigma2 node2 in
diff --git a/src/plugins/wp/tests/wp_hoare/oracle/refguards.res.oracle b/src/plugins/wp/tests/wp_hoare/oracle/refguards.res.oracle
index 86d8304e949ad0cad9c80264272451c1b50a699a..418bee4620c763c2a651cd882ab38b3da2d79a20 100644
--- a/src/plugins/wp/tests/wp_hoare/oracle/refguards.res.oracle
+++ b/src/plugins/wp/tests/wp_hoare/oracle/refguards.res.oracle
@@ -72,7 +72,6 @@ Prove: true.
 ------------------------------------------------------------
 
 Goal Post-condition 'KO' in 's':
-Assume { (* Heap *) Type: (region(c.base) <= 0) /\ (region(d.base) <= 0). }
 Prove: d != c.
 
 ------------------------------------------------------------
diff --git a/src/plugins/wp/tests/wp_usage/oracle/caveat2.res.oracle b/src/plugins/wp/tests/wp_usage/oracle/caveat2.res.oracle
index f6b9a283a93f75b56fb7bc4ebb5e73c6d67787a4..cdee7471627c5632a2428839fb657ac8e2c78328 100644
--- a/src/plugins/wp/tests/wp_usage/oracle/caveat2.res.oracle
+++ b/src/plugins/wp/tests/wp_usage/oracle/caveat2.res.oracle
@@ -25,11 +25,8 @@ Assume {
   When: (0 <= i) /\ (i < n).
   (* Pre-condition *)
   Have: n <= 3.
-  Have: ({ Init_p_0 with Init_F1_S_n = true }) = Init_p_0.
   (* Invariant *)
   Have: 0 <= n.
-  (* Loop assigns ... *)
-  Have: ({ Init_p_0 with Init_F1_S_a = v_1 }) = Init_p_0.
   (* Invariant *)
   Have: (0 <= i_1) /\ (i_1 <= n).
   (* Invariant *)
@@ -60,11 +57,8 @@ Assume {
   When: (i_1 <= i) /\ (0 <= i_1).
   (* Pre-condition *)
   Have: n <= 3.
-  Have: ({ Init_p_0 with Init_F1_S_n = true }) = Init_p_0.
   (* Invariant *)
   Have: 0 <= n.
-  (* Loop assigns ... *)
-  Have: ({ Init_p_0 with Init_F1_S_a = v_1 }) = Init_p_0.
   (* Invariant *)
   Have: (0 <= i) /\ (i <= n).
   (* Invariant *)
diff --git a/src/plugins/wp/wp.ml b/src/plugins/wp/wp.ml
index 0b8590d4dcdeb45062842ea3e71224b50f05e11c..9f420912bfd267be70a3a222dab082df5a0f87b2 100644
--- a/src/plugins/wp/wp.ml
+++ b/src/plugins/wp/wp.ml
@@ -127,7 +127,7 @@ module Layout = Layout
 
 (** {2 Compilers} *)
 
-module Sigs = Sigs
+module Memory = Memory
 module Driver = Driver
 module Context = Context
 module Ctypes = Ctypes
diff --git a/src/plugins/wp/wpPropId.ml b/src/plugins/wp/wpPropId.ml
index 4403bece1e5df182ef780db9c3d4afd8aeee24dd..8f1005da4edc668718f4112cf1870c7eb04dec42 100644
--- a/src/plugins/wp/wpPropId.ml
+++ b/src/plugins/wp/wpPropId.ml
@@ -50,7 +50,7 @@ type prop_kind =
   | PKSmoke       (** expected to fail *)
   | PKPre of kernel_function * stmt * Property.t
   (** precondition for function
-      at stmt, property of the require. Many information that should come
+      at stmt, property of the require. Mterm information that should come
       from the p_prop part of the prop_id, but in the PKPre case,
       it seems that it is hidden in a IPBlob property ! *)